src/ZF/List.thy
author wenzelm
Tue, 24 Sep 2024 21:31:20 +0200
changeset 80946 b76f64d7d493
parent 80788 66a8113ac23e
child 81011 6d34c2bedaa3
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68490
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 67443
diff changeset
     1
(*  Title:      ZF/List.thy
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     3
    Copyright   1994  University of Cambridge
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     4
*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
     6
section\<open>Lists in Zermelo-Fraenkel Set Theory\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     7
68490
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 67443
diff changeset
     8
theory List imports Datatype ArithSimp begin
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     9
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    10
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    11
  list       :: "i\<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    12
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    13
datatype
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    14
  "list(A)" = Nil | Cons ("a \<in> A", "l \<in> list(A)")
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    15
80787
f27f66e9adca more direct notation;
wenzelm
parents: 80761
diff changeset
    16
notation Nil (\<open>[]\<close>)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    17
80788
66a8113ac23e more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 80787
diff changeset
    18
nonterminal list_args
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    19
syntax
80788
66a8113ac23e more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 80787
diff changeset
    20
  "" :: "i \<Rightarrow> list_args"  (\<open>_\<close>)
66a8113ac23e more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 80787
diff changeset
    21
  "_List_args" :: "[i, list_args] \<Rightarrow> list_args"  (\<open>_,/ _\<close>)
66a8113ac23e more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 80787
diff changeset
    22
  "_List" :: "list_args \<Rightarrow> i"  (\<open>[(_)]\<close>)
66a8113ac23e more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 80787
diff changeset
    23
syntax_consts
66a8113ac23e more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 80787
diff changeset
    24
  "_List_args" "_List" \<rightleftharpoons> Cons
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    25
translations
35068
544867142ea4 modernized translations;
wenzelm
parents: 32960
diff changeset
    26
  "[x, xs]"     == "CONST Cons(x, [xs])"
544867142ea4 modernized translations;
wenzelm
parents: 32960
diff changeset
    27
  "[x]"         == "CONST Cons(x, [])"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    28
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    29
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    30
  length :: "i\<Rightarrow>i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    31
  hd     :: "i\<Rightarrow>i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    32
  tl     :: "i\<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    33
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    34
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    35
  "length([]) = 0"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    36
  "length(Cons(a,l)) = succ(length(l))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    37
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    38
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    39
  "hd([]) = 0"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    40
  "hd(Cons(a,l)) = a"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    41
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    42
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    43
  "tl([]) = []"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    44
  "tl(Cons(a,l)) = l"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    45
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    46
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    47
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    48
  map         :: "[i\<Rightarrow>i, i] \<Rightarrow> i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    49
  set_of_list :: "i\<Rightarrow>i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    50
  app         :: "[i,i]\<Rightarrow>i"                        (infixr \<open>@\<close> 60)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    51
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    52
(*map is a binding operator -- it applies to meta-level functions, not
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    53
object-level functions.  This simplifies the final form of term_rec_conv,
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    54
although complicating its derivation.*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    55
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    56
  "map(f,[]) = []"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    57
  "map(f,Cons(a,l)) = Cons(f(a), map(f,l))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    58
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    59
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    60
  "set_of_list([]) = 0"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    61
  "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    62
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    63
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    64
  app_Nil:  "[] @ ys = ys"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    65
  app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    66
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    67
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    68
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    69
  rev :: "i\<Rightarrow>i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    70
  flat       :: "i\<Rightarrow>i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    71
  list_add   :: "i\<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    72
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    73
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    74
  "rev([]) = []"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    75
  "rev(Cons(a,l)) = rev(l) @ [a]"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    76
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    77
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    78
  "flat([]) = []"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    79
  "flat(Cons(l,ls)) = l @ flat(ls)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    80
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    81
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    82
  "list_add([]) = 0"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    83
  "list_add(Cons(a,l)) = a #+ list_add(l)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    84
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    85
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    86
  drop       :: "[i,i]\<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    87
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    88
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    89
  drop_0:    "drop(0,l) = l"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    90
  drop_succ: "drop(succ(i), l) = tl (drop(i,l))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    91
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    92
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    93
(*** Thanks to Sidi Ehmety for the following ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    94
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    95
definition
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    96
(* Function `take' returns the first n elements of a list *)
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    97
  take     :: "[i,i]\<Rightarrow>i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
    98
  "take(n, as) \<equiv> list_rec(\<lambda>n\<in>nat. [],
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    99
                \<lambda>a l r. \<lambda>n\<in>nat. nat_case([], \<lambda>m. Cons(a, r`m), n), as)`n"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   100
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   101
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   102
  nth :: "[i, i]\<Rightarrow>i"  where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 61798
diff changeset
   103
  \<comment> \<open>returns the (n+1)th element of a list, or 0 if the
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   104
   list is too short.\<close>
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   105
  "nth(n, as) \<equiv> list_rec(\<lambda>n\<in>nat. 0,
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   106
                          \<lambda>a l r. \<lambda>n\<in>nat. nat_case(a, \<lambda>m. r`m, n), as) ` n"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   107
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   108
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   109
  list_update :: "[i, i, i]\<Rightarrow>i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   110
  "list_update(xs, i, v) \<equiv> list_rec(\<lambda>n\<in>nat. Nil,
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   111
      \<lambda>u us vs. \<lambda>n\<in>nat. nat_case(Cons(v, us), \<lambda>m. Cons(u, vs`m), n), xs)`i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   112
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   113
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   114
  filter :: "[i\<Rightarrow>o, i] \<Rightarrow> i"
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   115
  upt :: "[i, i] \<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   116
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   117
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   118
  "filter(P, Nil) = Nil"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   119
  "filter(P, Cons(x, xs)) =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   120
     (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   121
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   122
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   123
  "upt(i, 0) = Nil"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   124
  "upt(i, succ(j)) = (if i \<le> j then upt(i, j)@[j] else Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   125
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   126
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   127
  min :: "[i,i] \<Rightarrow>i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   128
    "min(x, y) \<equiv> (if x \<le> y then x else y)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   129
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   130
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   131
  max :: "[i, i] \<Rightarrow>i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   132
    "max(x, y) \<equiv> (if x \<le> y then y else x)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   133
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   134
(*** Aspects of the datatype definition ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   135
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   136
declare list.intros [simp,TC]
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   137
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   138
(*An elimination rule, for type-checking*)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   139
inductive_cases ConsE: "Cons(a,l) \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   140
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   141
lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) \<longleftrightarrow> a \<in> A \<and> l \<in> list(A)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   142
by (blast elim: ConsE)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   143
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   144
(*Proving freeness results*)
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   145
lemma Cons_iff: "Cons(a,l)=Cons(a',l') \<longleftrightarrow> a=a' \<and> l=l'"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   146
by auto
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   147
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   148
lemma Nil_Cons_iff: "\<not> Nil=Cons(a,l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   149
by auto
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   150
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   151
lemma list_unfold: "list(A) = {0} + (A * list(A))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   152
by (blast intro!: list.intros [unfolded list.con_defs]
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   153
          elim: list.cases [unfolded list.con_defs])
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   154
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   155
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   156
(**  Lemmas to justify using "list" in other recursive type definitions **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   157
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   158
lemma list_mono: "A<=B \<Longrightarrow> list(A) \<subseteq> list(B)"
76217
8655344f1cf6 More obsolete "unfold" calls
paulson <lp15@cam.ac.uk>
parents: 76216
diff changeset
   159
  unfolding list.defs 
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   160
apply (rule lfp_mono)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   161
apply (simp_all add: list.bnd_mono)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   162
apply (assumption | rule univ_mono basic_monos)+
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   163
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   164
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   165
(*There is a similar proof by list induction.*)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   166
lemma list_univ: "list(univ(A)) \<subseteq> univ(A)"
76217
8655344f1cf6 More obsolete "unfold" calls
paulson <lp15@cam.ac.uk>
parents: 76216
diff changeset
   167
  unfolding list.defs list.con_defs
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   168
apply (rule lfp_lowerbound)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   169
apply (rule_tac [2] A_subset_univ [THEN univ_mono])
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   170
apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   171
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   172
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   173
(*These two theorems justify datatypes involving list(nat), list(A), ...*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   174
lemmas list_subset_univ = subset_trans [OF list_mono list_univ]
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   175
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   176
lemma list_into_univ: "\<lbrakk>l \<in> list(A);  A \<subseteq> univ(B)\<rbrakk> \<Longrightarrow> l \<in> univ(B)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   177
by (blast intro: list_subset_univ [THEN subsetD])
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   178
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   179
lemma list_case_type:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   180
    "\<lbrakk>l \<in> list(A);
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   181
        c \<in> C(Nil);
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   182
        \<And>x y. \<lbrakk>x \<in> A;  y \<in> list(A)\<rbrakk> \<Longrightarrow> h(x,y): C(Cons(x,y))
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   183
\<rbrakk> \<Longrightarrow> list_case(c,h,l) \<in> C(l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   184
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   185
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   186
lemma list_0_triv: "list(0) = {Nil}"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   187
apply (rule equalityI, auto)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   188
apply (induct_tac x, auto)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   189
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   190
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   191
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   192
(*** List functions ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   193
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   194
lemma tl_type: "l \<in> list(A) \<Longrightarrow> tl(l) \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   195
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   196
apply (simp_all (no_asm_simp) add: list.intros)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   197
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   198
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   199
(** drop **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   200
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   201
lemma drop_Nil [simp]: "i \<in> nat \<Longrightarrow> drop(i, Nil) = Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   202
apply (induct_tac "i")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   203
apply (simp_all (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   204
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   205
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   206
lemma drop_succ_Cons [simp]: "i \<in> nat \<Longrightarrow> drop(succ(i), Cons(a,l)) = drop(i,l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   207
apply (rule sym)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   208
apply (induct_tac "i")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   209
apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   210
apply (simp (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   211
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   212
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   213
lemma drop_type [simp,TC]: "\<lbrakk>i \<in> nat; l \<in> list(A)\<rbrakk> \<Longrightarrow> drop(i,l) \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   214
apply (induct_tac "i")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   215
apply (simp_all (no_asm_simp) add: tl_type)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   216
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   217
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   218
declare drop_succ [simp del]
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   219
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   220
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   221
(** Type checking -- proved by induction, as usual **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   222
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   223
lemma list_rec_type [TC]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   224
    "\<lbrakk>l \<in> list(A);
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   225
        c \<in> C(Nil);
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   226
        \<And>x y r. \<lbrakk>x \<in> A;  y \<in> list(A);  r \<in> C(y)\<rbrakk> \<Longrightarrow> h(x,y,r): C(Cons(x,y))
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   227
\<rbrakk> \<Longrightarrow> list_rec(c,h,l) \<in> C(l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   228
by (induct_tac "l", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   229
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   230
(** map **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   231
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   232
lemma map_type [TC]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   233
    "\<lbrakk>l \<in> list(A);  \<And>x. x \<in> A \<Longrightarrow> h(x): B\<rbrakk> \<Longrightarrow> map(h,l) \<in> list(B)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   234
apply (simp add: map_list_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   235
apply (typecheck add: list.intros list_rec_type, blast)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   236
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   237
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   238
lemma map_type2 [TC]: "l \<in> list(A) \<Longrightarrow> map(h,l) \<in> list({h(u). u \<in> A})"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   239
apply (erule map_type)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   240
apply (erule RepFunI)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   241
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   242
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   243
(** length **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   244
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   245
lemma length_type [TC]: "l \<in> list(A) \<Longrightarrow> length(l) \<in> nat"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   246
by (simp add: length_list_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   247
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   248
lemma lt_length_in_nat:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   249
   "\<lbrakk>x < length(xs); xs \<in> list(A)\<rbrakk> \<Longrightarrow> x \<in> nat"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   250
by (frule lt_nat_in_nat, typecheck)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   251
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   252
(** app **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   253
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   254
lemma app_type [TC]: "\<lbrakk>xs: list(A);  ys: list(A)\<rbrakk> \<Longrightarrow> xs@ys \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   255
by (simp add: app_list_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   256
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   257
(** rev **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   258
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   259
lemma rev_type [TC]: "xs: list(A) \<Longrightarrow> rev(xs) \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   260
by (simp add: rev_list_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   261
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   262
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   263
(** flat **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   264
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   265
lemma flat_type [TC]: "ls: list(list(A)) \<Longrightarrow> flat(ls) \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   266
by (simp add: flat_list_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   267
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   268
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   269
(** set_of_list **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   270
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   271
lemma set_of_list_type [TC]: "l \<in> list(A) \<Longrightarrow> set_of_list(l) \<in> Pow(A)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   272
  unfolding set_of_list_list_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   273
apply (erule list_rec_type, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   274
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   275
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   276
lemma set_of_list_append:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   277
     "xs: list(A) \<Longrightarrow> set_of_list (xs@ys) = set_of_list(xs) \<union> set_of_list(ys)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   278
apply (erule list.induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   279
apply (simp_all (no_asm_simp) add: Un_cons)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   280
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   281
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   282
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   283
(** list_add **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   284
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   285
lemma list_add_type [TC]: "xs: list(nat) \<Longrightarrow> list_add(xs) \<in> nat"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   286
by (simp add: list_add_list_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   287
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   288
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   289
(*** theorems about map ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   290
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   291
lemma map_ident [simp]: "l \<in> list(A) \<Longrightarrow> map(\<lambda>u. u, l) = l"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   292
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   293
apply (simp_all (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   294
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   295
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   296
lemma map_compose: "l \<in> list(A) \<Longrightarrow> map(h, map(j,l)) = map(\<lambda>u. h(j(u)), l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   297
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   298
apply (simp_all (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   299
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   300
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   301
lemma map_app_distrib: "xs: list(A) \<Longrightarrow> map(h, xs@ys) = map(h,xs) @ map(h,ys)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   302
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   303
apply (simp_all (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   304
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   305
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   306
lemma map_flat: "ls: list(list(A)) \<Longrightarrow> map(h, flat(ls)) = flat(map(map(h),ls))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   307
apply (induct_tac "ls")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   308
apply (simp_all (no_asm_simp) add: map_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   309
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   310
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   311
lemma list_rec_map:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   312
     "l \<in> list(A) \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   313
      list_rec(c, d, map(h,l)) =
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   314
      list_rec(c, \<lambda>x xs r. d(h(x), map(h,xs), r), l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   315
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   316
apply (simp_all (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   317
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   318
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   319
(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   320
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   321
(* @{term"c \<in> list(Collect(B,P)) \<Longrightarrow> c \<in> list"} *)
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 35425
diff changeset
   322
lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   323
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   324
lemma map_list_Collect: "l \<in> list({x \<in> A. h(x)=j(x)}) \<Longrightarrow> map(h,l) = map(j,l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   325
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   326
apply (simp_all (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   327
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   328
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   329
(*** theorems about length ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   330
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   331
lemma length_map [simp]: "xs: list(A) \<Longrightarrow> length(map(h,xs)) = length(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   332
by (induct_tac "xs", simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   333
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   334
lemma length_app [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   335
     "\<lbrakk>xs: list(A); ys: list(A)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   336
      \<Longrightarrow> length(xs@ys) = length(xs) #+ length(ys)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   337
by (induct_tac "xs", simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   338
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   339
lemma length_rev [simp]: "xs: list(A) \<Longrightarrow> length(rev(xs)) = length(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   340
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   341
apply (simp_all (no_asm_simp) add: length_app)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   342
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   343
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   344
lemma length_flat:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   345
     "ls: list(list(A)) \<Longrightarrow> length(flat(ls)) = list_add(map(length,ls))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   346
apply (induct_tac "ls")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   347
apply (simp_all (no_asm_simp) add: length_app)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   348
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   349
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   350
(** Length and drop **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   351
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   352
(*Lemma for the inductive step of drop_length*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   353
lemma drop_length_Cons [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   354
     "xs: list(A) \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   355
           \<forall>x.  \<exists>z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   356
by (erule list.induct, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   357
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   358
lemma drop_length [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   359
     "l \<in> list(A) \<Longrightarrow> \<forall>i \<in> length(l). (\<exists>z zs. drop(i,l) = Cons(z,zs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   360
apply (erule list.induct, simp_all, safe)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   361
apply (erule drop_length_Cons)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   362
apply (rule natE)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   363
apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   364
apply (blast intro: succ_in_naturalD length_type)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   365
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   366
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   367
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   368
(*** theorems about app ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   369
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   370
lemma app_right_Nil [simp]: "xs: list(A) \<Longrightarrow> xs@Nil=xs"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   371
by (erule list.induct, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   372
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   373
lemma app_assoc: "xs: list(A) \<Longrightarrow> (xs@ys)@zs = xs@(ys@zs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   374
by (induct_tac "xs", simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   375
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   376
lemma flat_app_distrib: "ls: list(list(A)) \<Longrightarrow> flat(ls@ms) = flat(ls)@flat(ms)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   377
apply (induct_tac "ls")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   378
apply (simp_all (no_asm_simp) add: app_assoc)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   379
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   380
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   381
(*** theorems about rev ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   382
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   383
lemma rev_map_distrib: "l \<in> list(A) \<Longrightarrow> rev(map(h,l)) = map(h,rev(l))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   384
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   385
apply (simp_all (no_asm_simp) add: map_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   386
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   387
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   388
(*Simplifier needs the premises as assumptions because rewriting will not
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   389
  instantiate the variable ?A in the rules' typing conditions; note that
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   390
  rev_type does not instantiate ?A.  Only the premises do.
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   391
*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   392
lemma rev_app_distrib:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   393
     "\<lbrakk>xs: list(A);  ys: list(A)\<rbrakk> \<Longrightarrow> rev(xs@ys) = rev(ys)@rev(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   394
apply (erule list.induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   395
apply (simp_all add: app_assoc)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   396
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   397
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   398
lemma rev_rev_ident [simp]: "l \<in> list(A) \<Longrightarrow> rev(rev(l))=l"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   399
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   400
apply (simp_all (no_asm_simp) add: rev_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   401
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   402
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   403
lemma rev_flat: "ls: list(list(A)) \<Longrightarrow> rev(flat(ls)) = flat(map(rev,rev(ls)))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   404
apply (induct_tac "ls")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   405
apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   406
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   407
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   408
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   409
(*** theorems about list_add ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   410
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   411
lemma list_add_app:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   412
     "\<lbrakk>xs: list(nat);  ys: list(nat)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   413
      \<Longrightarrow> list_add(xs@ys) = list_add(ys) #+ list_add(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   414
apply (induct_tac "xs", simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   415
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   416
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   417
lemma list_add_rev: "l \<in> list(nat) \<Longrightarrow> list_add(rev(l)) = list_add(l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   418
apply (induct_tac "l")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   419
apply (simp_all (no_asm_simp) add: list_add_app)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   420
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   421
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   422
lemma list_add_flat:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   423
     "ls: list(list(nat)) \<Longrightarrow> list_add(flat(ls)) = list_add(map(list_add,ls))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   424
apply (induct_tac "ls")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   425
apply (simp_all (no_asm_simp) add: list_add_app)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   426
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   427
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   428
(** New induction rules **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   429
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   430
lemma list_append_induct [case_names Nil snoc, consumes 1]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   431
    "\<lbrakk>l \<in> list(A);
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   432
        P(Nil);
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   433
        \<And>x y. \<lbrakk>x \<in> A;  y \<in> list(A);  P(y)\<rbrakk> \<Longrightarrow> P(y @ [x])
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   434
\<rbrakk> \<Longrightarrow> P(l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   435
apply (subgoal_tac "P(rev(rev(l)))", simp)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   436
apply (erule rev_type [THEN list.induct], simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   437
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   438
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   439
lemma list_complete_induct_lemma [rule_format]:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   440
 assumes ih:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   441
    "\<And>l. \<lbrakk>l \<in> list(A);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   442
             \<forall>l' \<in> list(A). length(l') < length(l) \<longrightarrow> P(l')\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   443
          \<Longrightarrow> P(l)"
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   444
  shows "n \<in> nat \<Longrightarrow> \<forall>l \<in> list(A). length(l) < n \<longrightarrow> P(l)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   445
apply (induct_tac n, simp)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   446
apply (blast intro: ih elim!: leE)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   447
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   448
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   449
theorem list_complete_induct:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   450
      "\<lbrakk>l \<in> list(A);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   451
          \<And>l. \<lbrakk>l \<in> list(A);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   452
                  \<forall>l' \<in> list(A). length(l') < length(l) \<longrightarrow> P(l')\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   453
               \<Longrightarrow> P(l)
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   454
\<rbrakk> \<Longrightarrow> P(l)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   455
apply (rule list_complete_induct_lemma [of A])
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   456
   prefer 4 apply (rule le_refl, simp)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   457
  apply blast
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   458
 apply simp
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   459
apply assumption
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   460
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   461
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   462
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   463
(*** Thanks to Sidi Ehmety for these results about min, take, etc. ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   464
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   465
(** min FIXME: replace by Int! **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   466
(* Min theorems are also true for i, j ordinals *)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   467
lemma min_sym: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> min(i,j)=min(j,i)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   468
  unfolding min_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   469
apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   470
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   471
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   472
lemma min_type [simp,TC]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> min(i,j):nat"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   473
by (unfold min_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   474
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   475
lemma min_0 [simp]: "i \<in> nat \<Longrightarrow> min(0,i) = 0"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   476
  unfolding min_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   477
apply (auto dest: not_lt_imp_le)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   478
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   479
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   480
lemma min_02 [simp]: "i \<in> nat \<Longrightarrow> min(i, 0) = 0"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   481
  unfolding min_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   482
apply (auto dest: not_lt_imp_le)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   483
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   484
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   485
lemma lt_min_iff: "\<lbrakk>i \<in> nat; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> i<min(j,k) \<longleftrightarrow> i<j \<and> i<k"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   486
  unfolding min_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   487
apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   488
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   489
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   490
lemma min_succ_succ [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   491
     "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow>  min(succ(i), succ(j))= succ(min(i, j))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   492
apply (unfold min_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   493
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   494
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   495
(*** more theorems about lists ***)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   496
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   497
(** filter **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   498
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   499
lemma filter_append [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   500
     "xs:list(A) \<Longrightarrow> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   501
by (induct_tac "xs", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   502
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   503
lemma filter_type [simp,TC]: "xs:list(A) \<Longrightarrow> filter(P, xs):list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   504
by (induct_tac "xs", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   505
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   506
lemma length_filter: "xs:list(A) \<Longrightarrow> length(filter(P, xs)) \<le> length(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   507
apply (induct_tac "xs", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   508
apply (rule_tac j = "length (l) " in le_trans)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   509
apply (auto simp add: le_iff)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   510
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   511
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   512
lemma filter_is_subset: "xs:list(A) \<Longrightarrow> set_of_list(filter(P,xs)) \<subseteq> set_of_list(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   513
by (induct_tac "xs", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   514
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   515
lemma filter_False [simp]: "xs:list(A) \<Longrightarrow> filter(\<lambda>p. False, xs) = Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   516
by (induct_tac "xs", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   517
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   518
lemma filter_True [simp]: "xs:list(A) \<Longrightarrow> filter(\<lambda>p. True, xs) = xs"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   519
by (induct_tac "xs", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   520
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   521
(** length **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   522
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   523
lemma length_is_0_iff [simp]: "xs:list(A) \<Longrightarrow> length(xs)=0 \<longleftrightarrow> xs=Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   524
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   525
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   526
lemma length_is_0_iff2 [simp]: "xs:list(A) \<Longrightarrow> 0 = length(xs) \<longleftrightarrow> xs=Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   527
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   528
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   529
lemma length_tl [simp]: "xs:list(A) \<Longrightarrow> length(tl(xs)) = length(xs) #- 1"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   530
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   531
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   532
lemma length_greater_0_iff: "xs:list(A) \<Longrightarrow> 0<length(xs) \<longleftrightarrow> xs \<noteq> Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   533
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   534
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   535
lemma length_succ_iff: "xs:list(A) \<Longrightarrow> length(xs)=succ(n) \<longleftrightarrow> (\<exists>y ys. xs=Cons(y, ys) \<and> length(ys)=n)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   536
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   537
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   538
(** more theorems about append **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   539
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   540
lemma append_is_Nil_iff [simp]:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   541
     "xs:list(A) \<Longrightarrow> (xs@ys = Nil) \<longleftrightarrow> (xs=Nil \<and> ys = Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   542
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   543
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   544
lemma append_is_Nil_iff2 [simp]:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   545
     "xs:list(A) \<Longrightarrow> (Nil = xs@ys) \<longleftrightarrow> (xs=Nil \<and> ys = Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   546
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   547
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   548
lemma append_left_is_self_iff [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   549
     "xs:list(A) \<Longrightarrow> (xs@ys = xs) \<longleftrightarrow> (ys = Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   550
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   551
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   552
lemma append_left_is_self_iff2 [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   553
     "xs:list(A) \<Longrightarrow> (xs = xs@ys) \<longleftrightarrow> (ys = Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   554
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   555
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   556
(*TOO SLOW as a default simprule!*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   557
lemma append_left_is_Nil_iff [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   558
     "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow>
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   559
   length(ys)=length(zs) \<longrightarrow> (xs@ys=zs \<longleftrightarrow> (xs=Nil \<and> ys=zs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   560
apply (erule list.induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   561
apply (auto simp add: length_app)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   562
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   563
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   564
(*TOO SLOW as a default simprule!*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   565
lemma append_left_is_Nil_iff2 [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   566
     "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow>
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   567
   length(ys)=length(zs) \<longrightarrow> (zs=ys@xs \<longleftrightarrow> (xs=Nil \<and> ys=zs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   568
apply (erule list.induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   569
apply (auto simp add: length_app)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   570
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   571
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   572
lemma append_eq_append_iff [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   573
     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A).
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   574
      length(xs)=length(ys) \<longrightarrow> (xs@us = ys@vs) \<longleftrightarrow> (xs=ys \<and> us=vs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   575
apply (erule list.induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   576
apply (simp (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   577
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   578
apply (erule_tac a = ys in list.cases, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   579
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   580
declare append_eq_append_iff [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   581
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   582
lemma append_eq_append [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   583
  "xs:list(A) \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   584
   \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   585
   length(us) = length(vs) \<longrightarrow> (xs@us = ys@vs) \<longrightarrow> (xs=ys \<and> us=vs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   586
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   587
apply (force simp add: length_app, clarify)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   588
apply (erule_tac a = ys in list.cases, simp)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   589
apply (subgoal_tac "Cons (a, l) @ us =vs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   590
 apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   591
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   592
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   593
lemma append_eq_append_iff2 [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   594
 "\<lbrakk>xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs)\<rbrakk>
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   595
  \<Longrightarrow>  xs@us = ys@vs \<longleftrightarrow> (xs=ys \<and> us=vs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   596
apply (rule iffI)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   597
apply (rule append_eq_append, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   598
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   599
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   600
lemma append_self_iff [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   601
     "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow> xs@ys=xs@zs \<longleftrightarrow> ys=zs"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   602
by simp
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   603
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   604
lemma append_self_iff2 [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   605
     "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow> ys@xs=zs@xs \<longleftrightarrow> ys=zs"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   606
by simp
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   607
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   608
(* Can also be proved from append_eq_append_iff2,
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   609
but the proof requires two more hypotheses: x \<in> A and y \<in> A *)
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   610
lemma append1_eq_iff [rule_format]:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   611
     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A). xs@[x] = ys@[y] \<longleftrightarrow> (xs = ys \<and> x=y)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   612
apply (erule list.induct)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   613
 apply clarify
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   614
 apply (erule list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   615
 apply simp_all
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   616
txt\<open>Inductive step\<close>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   617
apply clarify
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   618
apply (erule_tac a=ys in list.cases, simp_all)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   619
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   620
declare append1_eq_iff [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   621
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   622
lemma append_right_is_self_iff [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   623
     "\<lbrakk>xs:list(A); ys:list(A)\<rbrakk> \<Longrightarrow> (xs@ys = ys) \<longleftrightarrow> (xs=Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   624
by (simp (no_asm_simp) add: append_left_is_Nil_iff)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   625
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   626
lemma append_right_is_self_iff2 [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   627
     "\<lbrakk>xs:list(A); ys:list(A)\<rbrakk> \<Longrightarrow> (ys = xs@ys) \<longleftrightarrow> (xs=Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   628
apply (rule iffI)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   629
apply (drule sym, auto)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   630
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   631
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   632
lemma hd_append [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   633
     "xs:list(A) \<Longrightarrow> xs \<noteq> Nil \<longrightarrow> hd(xs @ ys) = hd(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   634
by (induct_tac "xs", auto)
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   635
declare hd_append [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   636
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   637
lemma tl_append [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   638
     "xs:list(A) \<Longrightarrow> xs\<noteq>Nil \<longrightarrow> tl(xs @ ys) = tl(xs)@ys"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   639
by (induct_tac "xs", auto)
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   640
declare tl_append [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   641
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   642
(** rev **)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   643
lemma rev_is_Nil_iff [simp]: "xs:list(A) \<Longrightarrow> (rev(xs) = Nil \<longleftrightarrow> xs = Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   644
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   645
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   646
lemma Nil_is_rev_iff [simp]: "xs:list(A) \<Longrightarrow> (Nil = rev(xs) \<longleftrightarrow> xs = Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   647
by (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   648
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   649
lemma rev_is_rev_iff [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   650
     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A). rev(xs)=rev(ys) \<longleftrightarrow> xs=ys"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   651
apply (erule list.induct, force, clarify)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   652
apply (erule_tac a = ys in list.cases, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   653
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   654
declare rev_is_rev_iff [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   655
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   656
lemma rev_list_elim [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   657
     "xs:list(A) \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   658
      (xs=Nil \<longrightarrow> P) \<longrightarrow> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] \<longrightarrow>P)\<longrightarrow>P"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   659
by (erule list_append_induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   660
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   661
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   662
(** more theorems about drop **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   663
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   664
lemma length_drop [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   665
     "n \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). length(drop(n, xs)) = length(xs) #- n"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   666
apply (erule nat_induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   667
apply (auto elim: list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   668
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   669
declare length_drop [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   670
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   671
lemma drop_all [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   672
     "n \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> drop(n, xs)=Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   673
apply (erule nat_induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   674
apply (auto elim: list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   675
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   676
declare drop_all [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   677
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   678
lemma drop_append [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   679
     "n \<in> nat \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   680
      \<forall>xs \<in> list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   681
apply (induct_tac "n")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   682
apply (auto elim: list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   683
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   684
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   685
lemma drop_drop:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   686
    "m \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). \<forall>n \<in> nat. drop(n, drop(m, xs))=drop(n #+ m, xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   687
apply (induct_tac "m")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   688
apply (auto elim: list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   689
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   690
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   691
(** take **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   692
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   693
lemma take_0 [simp]: "xs:list(A) \<Longrightarrow> take(0, xs) =  Nil"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   694
  unfolding take_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   695
apply (erule list.induct, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   696
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   697
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   698
lemma take_succ_Cons [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   699
    "n \<in> nat \<Longrightarrow> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   700
by (simp add: take_def)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   701
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   702
(* Needed for proving take_all *)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   703
lemma take_Nil [simp]: "n \<in> nat \<Longrightarrow> take(n, Nil) = Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   704
by (unfold take_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   705
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   706
lemma take_all [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   707
     "n \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). length(xs) \<le> n  \<longrightarrow> take(n, xs) = xs"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   708
apply (erule nat_induct)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   709
apply (auto elim: list.cases)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   710
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   711
declare take_all [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   712
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   713
lemma take_type [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   714
     "xs:list(A) \<Longrightarrow> \<forall>n \<in> nat. take(n, xs):list(A)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   715
apply (erule list.induct, simp, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   716
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   717
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   718
declare take_type [simp,TC]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   719
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   720
lemma take_append [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   721
 "xs:list(A) \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   722
  \<forall>ys \<in> list(A). \<forall>n \<in> nat. take(n, xs @ ys) =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   723
                            take(n, xs) @ take(n #- length(xs), ys)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   724
apply (erule list.induct, simp, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   725
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   726
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   727
declare take_append [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   728
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   729
lemma take_take [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   730
   "m \<in> nat \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   731
    \<forall>xs \<in> list(A). \<forall>n \<in> nat. take(n, take(m,xs))= take(min(n, m), xs)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   732
apply (induct_tac "m", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   733
apply (erule_tac a = xs in list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   734
apply (auto simp add: take_Nil)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   735
apply (erule_tac n=n in natE)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   736
apply (auto intro: take_0 take_type)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   737
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   738
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   739
(** nth **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   740
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   741
lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   742
by (simp add: nth_def)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   743
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   744
lemma nth_Cons [simp]: "n \<in> nat \<Longrightarrow> nth(succ(n), Cons(a,l)) = nth(n,l)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   745
by (simp add: nth_def)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   746
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   747
lemma nth_empty [simp]: "nth(n, Nil) = 0"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   748
by (simp add: nth_def)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   749
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   750
lemma nth_type [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   751
     "xs:list(A) \<Longrightarrow> \<forall>n. n < length(xs) \<longrightarrow> nth(n,xs) \<in> A"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   752
apply (erule list.induct, simp, clarify)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   753
apply (subgoal_tac "n \<in> nat")
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   754
 apply (erule natE, auto dest!: le_in_nat)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   755
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   756
declare nth_type [simp,TC]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   757
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   758
lemma nth_eq_0 [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   759
     "xs:list(A) \<Longrightarrow> \<forall>n \<in> nat. length(xs) \<le> n \<longrightarrow> nth(n,xs) = 0"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   760
apply (erule list.induct, simp, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   761
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   762
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   763
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   764
lemma nth_append [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   765
  "xs:list(A) \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   766
   \<forall>n \<in> nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   767
                                else nth(n #- length(xs), ys))"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   768
apply (induct_tac "xs", simp, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   769
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   770
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   771
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   772
lemma set_of_list_conv_nth:
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   773
    "xs:list(A)
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   774
     \<Longrightarrow> set_of_list(xs) = {x \<in> A. \<exists>i\<in>nat. i<length(xs) \<and> x = nth(i,xs)}"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   775
apply (induct_tac "xs", simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   776
apply (rule equalityI, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   777
apply (rule_tac x = 0 in bexI, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   778
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   779
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   780
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   781
(* Other theorems about lists *)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   782
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   783
lemma nth_take_lemma [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   784
 "k \<in> nat \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   785
  \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k \<le> length(xs) \<longrightarrow> k \<le> length(ys) \<longrightarrow>
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   786
      (\<forall>i \<in> nat. i<k \<longrightarrow> nth(i,xs) = nth(i,ys))\<longrightarrow> take(k,xs) = take(k,ys))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   787
apply (induct_tac "k")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   788
apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   789
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   790
(*Both lists are non-empty*)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   791
apply (erule_tac a=xs in list.cases, simp)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   792
apply (erule_tac a=ys in list.cases, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   793
apply (simp (no_asm_use) )
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   794
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   795
apply (simp (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   796
apply (rule conjI, force)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   797
apply (rename_tac y ys z zs)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   798
apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   799
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   800
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   801
lemma nth_equalityI [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   802
     "\<lbrakk>xs:list(A); ys:list(A); length(xs) = length(ys);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   803
         \<forall>i \<in> nat. i < length(xs) \<longrightarrow> nth(i,xs) = nth(i,ys)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   804
      \<Longrightarrow> xs = ys"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   805
apply (subgoal_tac "length (xs) \<le> length (ys) ")
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   806
apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   807
apply (simp_all add: take_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   808
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   809
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   810
(*The famous take-lemma*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   811
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   812
lemma take_equalityI [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   813
    "\<lbrakk>xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys))\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   814
     \<Longrightarrow> xs = ys"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   815
apply (case_tac "length (xs) \<le> length (ys) ")
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   816
apply (drule_tac x = "length (ys) " in bspec)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   817
apply (drule_tac [3] not_lt_imp_le)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   818
apply (subgoal_tac [5] "length (ys) \<le> length (xs) ")
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   819
apply (rule_tac [6] j = "succ (length (ys))" in le_trans)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   820
apply (rule_tac [6] leI)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   821
apply (drule_tac [5] x = "length (xs) " in bspec)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   822
apply (simp_all add: take_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   823
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   824
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   825
lemma nth_drop [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   826
  "n \<in> nat \<Longrightarrow> \<forall>i \<in> nat. \<forall>xs \<in> list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   827
apply (induct_tac "n", simp_all, clarify)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   828
apply (erule list.cases, auto)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   829
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   830
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   831
lemma take_succ [rule_format]:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   832
  "xs\<in>list(A)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   833
   \<Longrightarrow> \<forall>i. i < length(xs) \<longrightarrow> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   834
apply (induct_tac "xs", auto)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   835
apply (subgoal_tac "i\<in>nat")
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   836
apply (erule natE)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   837
apply (auto simp add: le_in_nat)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   838
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   839
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   840
lemma take_add [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   841
     "\<lbrakk>xs\<in>list(A); j\<in>nat\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   842
      \<Longrightarrow> \<forall>i\<in>nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   843
apply (induct_tac "xs", simp_all, clarify)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   844
apply (erule_tac n = i in natE, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   845
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   846
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   847
lemma length_take:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   848
     "l\<in>list(A) \<Longrightarrow> \<forall>n\<in>nat. length(take(n,l)) = min(n, length(l))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   849
apply (induct_tac "l", safe, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   850
apply (erule natE, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   851
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   852
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   853
subsection\<open>The function zip\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   854
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   855
text\<open>Crafty definition to eliminate a type argument\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   856
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   857
consts
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   858
  zip_aux        :: "[i,i]\<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   859
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   860
primrec (*explicit lambda is required because both arguments of "un" vary*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   861
  "zip_aux(B,[]) =
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   862
     (\<lambda>ys \<in> list(B). list_case([], \<lambda>y l. [], ys))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   863
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   864
  "zip_aux(B,Cons(x,l)) =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   865
     (\<lambda>ys \<in> list(B).
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   866
        list_case(Nil, \<lambda>y zs. Cons(\<langle>x,y\<rangle>, zip_aux(B,l)`zs), ys))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   867
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   868
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   869
  zip :: "[i, i]\<Rightarrow>i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   870
   "zip(xs, ys) \<equiv> zip_aux(set_of_list(ys),xs)`ys"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   871
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   872
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   873
(* zip equations *)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   874
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   875
lemma list_on_set_of_list: "xs \<in> list(A) \<Longrightarrow> xs \<in> list(set_of_list(xs))"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   876
apply (induct_tac xs, simp_all)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   877
apply (blast intro: list_mono [THEN subsetD])
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   878
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   879
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   880
lemma zip_Nil [simp]: "ys:list(A) \<Longrightarrow> zip(Nil, ys)=Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   881
apply (simp add: zip_def list_on_set_of_list [of _ A])
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   882
apply (erule list.cases, simp_all)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   883
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   884
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   885
lemma zip_Nil2 [simp]: "xs:list(A) \<Longrightarrow> zip(xs, Nil)=Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   886
apply (simp add: zip_def list_on_set_of_list [of _ A])
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   887
apply (erule list.cases, simp_all)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   888
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   889
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   890
lemma zip_aux_unique [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   891
     "\<lbrakk>B<=C;  xs \<in> list(A)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   892
      \<Longrightarrow> \<forall>ys \<in> list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   893
apply (induct_tac xs)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   894
 apply simp_all
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   895
 apply (blast intro: list_mono [THEN subsetD], clarify)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   896
apply (erule_tac a=ys in list.cases, auto)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   897
apply (blast intro: list_mono [THEN subsetD])
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   898
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   899
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   900
lemma zip_Cons_Cons [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   901
     "\<lbrakk>xs:list(A); ys:list(B); x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow>
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   902
      zip(Cons(x,xs), Cons(y, ys)) = Cons(\<langle>x,y\<rangle>, zip(xs, ys))"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   903
apply (simp add: zip_def, auto)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   904
apply (rule zip_aux_unique, auto)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   905
apply (simp add: list_on_set_of_list [of _ B])
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   906
apply (blast intro: list_on_set_of_list list_mono [THEN subsetD])
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   907
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   908
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   909
lemma zip_type [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   910
     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(B). zip(xs, ys):list(A*B)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   911
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   912
apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   913
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   914
apply (erule_tac a = ys in list.cases, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   915
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   916
declare zip_type [simp,TC]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   917
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   918
(* zip length *)
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   919
lemma length_zip [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   920
     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(B). length(zip(xs,ys)) =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   921
                                     min(length(xs), length(ys))"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   922
  unfolding min_def
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   923
apply (induct_tac "xs", simp_all, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   924
apply (erule_tac a = ys in list.cases, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   925
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   926
declare length_zip [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   927
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   928
lemma zip_append1 [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   929
 "\<lbrakk>ys:list(A); zs:list(B)\<rbrakk> \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   930
  \<forall>xs \<in> list(A). zip(xs @ ys, zs) =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   931
                 zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   932
apply (induct_tac "zs", force, clarify)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   933
apply (erule_tac a = xs in list.cases, simp_all)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   934
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   935
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   936
lemma zip_append2 [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   937
 "\<lbrakk>xs:list(A); zs:list(B)\<rbrakk> \<Longrightarrow> \<forall>ys \<in> list(B). zip(xs, ys@zs) =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   938
       zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   939
apply (induct_tac "xs", force, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   940
apply (erule_tac a = ys in list.cases, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   941
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   942
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   943
lemma zip_append [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   944
 "\<lbrakk>length(xs) = length(us); length(ys) = length(vs);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   945
     xs:list(A); us:list(B); ys:list(A); vs:list(B)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   946
  \<Longrightarrow> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   947
by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   948
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   949
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   950
lemma zip_rev [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   951
 "ys:list(B) \<Longrightarrow> \<forall>xs \<in> list(A).
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   952
    length(xs) = length(ys) \<longrightarrow> zip(rev(xs), rev(ys)) = rev(zip(xs, ys))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   953
apply (induct_tac "ys", force, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   954
apply (erule_tac a = xs in list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   955
apply (auto simp add: length_rev)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   956
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   957
declare zip_rev [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   958
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   959
lemma nth_zip [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   960
   "ys:list(B) \<Longrightarrow> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   961
                    i < length(xs) \<longrightarrow> i < length(ys) \<longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   962
                    nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, ys)>"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   963
apply (induct_tac "ys", force, clarify)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   964
apply (erule_tac a = xs in list.cases, simp)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   965
apply (auto elim: natE)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   966
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   967
declare nth_zip [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   968
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   969
lemma set_of_list_zip [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   970
     "\<lbrakk>xs:list(A); ys:list(B); i \<in> nat\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   971
      \<Longrightarrow> set_of_list(zip(xs, ys)) =
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   972
          {\<langle>x, y\<rangle>:A*B. \<exists>i\<in>nat. i < min(length(xs), length(ys))
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   973
          \<and> x = nth(i, xs) \<and> y = nth(i, ys)}"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   974
by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   975
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   976
(** list_update **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   977
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   978
lemma list_update_Nil [simp]: "i \<in> nat \<Longrightarrow>list_update(Nil, i, v) = Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   979
by (unfold list_update_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   980
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   981
lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   982
by (unfold list_update_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   983
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   984
lemma list_update_Cons_succ [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   985
  "n \<in> nat \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   986
    list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   987
apply (unfold list_update_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   988
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   989
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   990
lemma list_update_type [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
   991
     "\<lbrakk>xs:list(A); v \<in> A\<rbrakk> \<Longrightarrow> \<forall>n \<in> nat. list_update(xs, n, v):list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   992
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   993
apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   994
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   995
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   996
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   997
declare list_update_type [simp,TC]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   998
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
   999
lemma length_list_update [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1000
     "xs:list(A) \<Longrightarrow> \<forall>i \<in> nat. length(list_update(xs, i, v))=length(xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1001
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1002
apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1003
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1004
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1005
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1006
declare length_list_update [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1007
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1008
lemma nth_list_update [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1009
     "\<lbrakk>xs:list(A)\<rbrakk> \<Longrightarrow> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs)  \<longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1010
         nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1011
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1012
 apply simp_all
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1013
apply clarify
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1014
apply (rename_tac i j)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1015
apply (erule_tac n=i in natE)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1016
apply (erule_tac [2] n=j in natE)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1017
apply (erule_tac n=j in natE, simp_all, force)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1018
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1019
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1020
lemma nth_list_update_eq [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1021
     "\<lbrakk>i < length(xs); xs:list(A)\<rbrakk> \<Longrightarrow> nth(i, list_update(xs, i,x)) = x"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1022
by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1023
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1024
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1025
lemma nth_list_update_neq [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1026
  "xs:list(A) \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1027
     \<forall>i \<in> nat. \<forall>j \<in> nat. i \<noteq> j \<longrightarrow> nth(j, list_update(xs,i,x)) = nth(j,xs)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1028
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1029
 apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1030
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1031
apply (erule natE)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1032
apply (erule_tac [2] natE, simp_all)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1033
apply (erule natE, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1034
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1035
declare nth_list_update_neq [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1036
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1037
lemma list_update_overwrite [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1038
     "xs:list(A) \<Longrightarrow> \<forall>i \<in> nat. i < length(xs)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1039
   \<longrightarrow> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1040
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1041
 apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1042
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1043
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1044
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1045
declare list_update_overwrite [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1046
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1047
lemma list_update_same_conv [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1048
     "xs:list(A) \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1049
      \<forall>i \<in> nat. i < length(xs) \<longrightarrow>
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
  1050
                 (list_update(xs, i, x) = xs) \<longleftrightarrow> (nth(i, xs) = x)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1051
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1052
 apply (simp (no_asm))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1053
apply clarify
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1054
apply (erule natE, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1055
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1057
lemma update_zip [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1058
     "ys:list(B) \<Longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1059
      \<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> list(A).
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1060
        length(xs) = length(ys) \<longrightarrow>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1061
        list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)),
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1062
                                              list_update(ys, i, snd(xy)))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1063
apply (induct_tac "ys")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1064
 apply auto
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1065
apply (erule_tac a = xs in list.cases)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1066
apply (auto elim: natE)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1067
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1068
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1069
lemma set_update_subset_cons [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1070
  "xs:list(A) \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1071
   \<forall>i \<in> nat. set_of_list(list_update(xs, i, x)) \<subseteq> cons(x, set_of_list(xs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1072
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1073
 apply simp
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1074
apply (rule ballI)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1075
apply (erule natE, simp_all, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1076
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1077
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1078
lemma set_of_list_update_subsetI:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1079
     "\<lbrakk>set_of_list(xs) \<subseteq> A; xs:list(A); x \<in> A; i \<in> nat\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1080
   \<Longrightarrow> set_of_list(list_update(xs, i,x)) \<subseteq> A"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1081
apply (rule subset_trans)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1082
apply (rule set_update_subset_cons, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1083
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1084
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1085
(** upt **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1086
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1087
lemma upt_rec:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1088
     "j \<in> nat \<Longrightarrow> upt(i,j) = (if i<j then Cons(i, upt(succ(i), j)) else Nil)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1089
apply (induct_tac "j", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1090
apply (drule not_lt_imp_le)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1091
apply (auto simp: lt_Ord intro: le_anti_sym)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1092
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1093
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1094
lemma upt_conv_Nil [simp]: "\<lbrakk>j \<le> i; j \<in> nat\<rbrakk> \<Longrightarrow> upt(i,j) = Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1095
apply (subst upt_rec, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1096
apply (auto simp add: le_iff)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1097
apply (drule lt_asym [THEN notE], auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1098
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1099
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1100
(*Only needed if upt_Suc is deleted from the simpset*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1101
lemma upt_succ_append:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1102
     "\<lbrakk>i \<le> j; j \<in> nat\<rbrakk> \<Longrightarrow> upt(i,succ(j)) = upt(i, j)@[j]"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1103
by simp
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1104
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1105
lemma upt_conv_Cons:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1106
     "\<lbrakk>i<j; j \<in> nat\<rbrakk>  \<Longrightarrow> upt(i,j) = Cons(i,upt(succ(i),j))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1107
apply (rule trans)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1108
apply (rule upt_rec, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1109
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1110
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1111
lemma upt_type [simp,TC]: "j \<in> nat \<Longrightarrow> upt(i,j):list(nat)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1112
by (induct_tac "j", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1113
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1114
(*LOOPS as a simprule, since j<=j*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1115
lemma upt_add_eq_append:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1116
     "\<lbrakk>i \<le> j; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1117
apply (induct_tac "k")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1118
apply (auto simp add: app_assoc app_type)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1119
apply (rule_tac j = j in le_trans, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1120
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1121
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1122
lemma length_upt [simp]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow>length(upt(i,j)) = j #- i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1123
apply (induct_tac "j")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1124
apply (rule_tac [2] sym)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1125
apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1126
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1127
71089
907b7a6471a0 tuned proofs;
wenzelm
parents: 71082
diff changeset
  1128
lemma nth_upt [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1129
     "\<lbrakk>i \<in> nat; j \<in> nat; k \<in> nat; i #+ k < j\<rbrakk> \<Longrightarrow> nth(k, upt(i,j)) = i #+ k"
71089
907b7a6471a0 tuned proofs;
wenzelm
parents: 71082
diff changeset
  1130
apply (rotate_tac -1, erule rev_mp)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1131
apply (induct_tac "j", simp)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1132
apply (auto dest!: not_lt_imp_le
71089
907b7a6471a0 tuned proofs;
wenzelm
parents: 71082
diff changeset
  1133
            simp add: nth_append le_iff less_diff_conv add_commute)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1134
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1135
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1136
lemma take_upt [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1137
     "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1138
         \<forall>i \<in> nat. i #+ m \<le> n \<longrightarrow> take(m, upt(i,n)) = upt(i,i#+m)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1139
apply (induct_tac "m")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1140
apply (simp (no_asm_simp) add: take_0)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1141
apply clarify
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1142
apply (subst upt_rec, simp)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1143
apply (rule sym)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1144
apply (subst upt_rec, simp)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1145
apply (simp_all del: upt.simps)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1146
apply (rule_tac j = "succ (i #+ x) " in lt_trans2)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1147
apply auto
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1148
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1149
declare take_upt [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1150
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1151
lemma map_succ_upt:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1152
     "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> map(succ, upt(m,n))= upt(succ(m), succ(n))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1153
apply (induct_tac "n")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1154
apply (auto simp add: map_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1155
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1156
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1157
lemma nth_map [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1158
     "xs:list(A) \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1159
      \<forall>n \<in> nat. n < length(xs) \<longrightarrow> nth(n, map(f, xs)) = f(nth(n, xs))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1160
apply (induct_tac "xs", simp)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1161
apply (rule ballI)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1162
apply (induct_tac "n", auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1163
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1164
declare nth_map [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1165
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1166
lemma nth_map_upt [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1167
     "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1168
      \<forall>i \<in> nat. i < n #- m \<longrightarrow> nth(i, map(f, upt(m,n))) = f(m #+ i)"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1169
apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1170
apply (subst map_succ_upt [symmetric], simp_all, clarify)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1171
apply (subgoal_tac "i < length (upt (0, x))")
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1172
 prefer 2
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1173
 apply (simp add: less_diff_conv)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1174
 apply (rule_tac j = "succ (i #+ y) " in lt_trans2)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1175
  apply simp
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1176
 apply simp
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1177
apply (subgoal_tac "i < length (upt (y, x))")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1178
 apply (simp_all add: add_commute less_diff_conv)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1179
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1180
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1181
(** sublist (a generalization of nth to sets) **)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1182
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1183
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
  1184
  sublist :: "[i, i] \<Rightarrow> i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1185
    "sublist(xs, A)\<equiv>
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
  1186
     map(fst, (filter(\<lambda>p. snd(p): A, zip(xs, upt(0,length(xs))))))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1187
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1188
lemma sublist_0 [simp]: "xs:list(A) \<Longrightarrow>sublist(xs, 0) =Nil"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1189
by (unfold sublist_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1190
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1191
lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1192
by (unfold sublist_def, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1193
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1194
lemma sublist_shift_lemma:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1195
 "\<lbrakk>xs:list(B); i \<in> nat\<rbrakk> \<Longrightarrow>
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
  1196
  map(fst, filter(\<lambda>p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
  1197
  map(fst, filter(\<lambda>p. snd(p):nat \<and> snd(p) #+ i \<in> A, zip(xs,upt(0,length(xs)))))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1198
apply (erule list_append_induct)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1199
apply (simp (no_asm_simp))
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1200
apply (auto simp add: add_commute length_app filter_append map_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1201
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1202
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1203
lemma sublist_type [simp,TC]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1204
     "xs:list(B) \<Longrightarrow> sublist(xs, A):list(B)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
  1205
  unfolding sublist_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1206
apply (induct_tac "xs")
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1207
apply (auto simp add: filter_append map_app_distrib)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1208
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1209
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1210
lemma upt_add_eq_append2:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1211
     "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1212
by (simp add: upt_add_eq_append [of 0] nat_0_le)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1213
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1214
lemma sublist_append:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1215
 "\<lbrakk>xs:list(B); ys:list(B)\<rbrakk> \<Longrightarrow>
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
  1216
  sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \<in> nat. j #+ length(xs): A})"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
  1217
  unfolding sublist_def
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1218
apply (erule_tac l = ys in list_append_induct, simp)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1219
apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric])
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1220
apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1221
apply (simp_all add: add_commute)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1222
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1223
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1224
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1225
lemma sublist_Cons:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1226
     "\<lbrakk>xs:list(B); x \<in> B\<rbrakk> \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1227
      sublist(Cons(x, xs), A) =
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
  1228
      (if 0 \<in> A then [x] else []) @ sublist(xs, {j \<in> nat. succ(j) \<in> A})"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1229
apply (erule_tac l = xs in list_append_induct)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1230
apply (simp (no_asm_simp) add: sublist_def)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1231
apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1232
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1233
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1234
lemma sublist_singleton [simp]:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1235
     "sublist([x], A) = (if 0 \<in> A then [x] else [])"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1236
by (simp add: sublist_Cons)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1237
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1238
lemma sublist_upt_eq_take [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1239
    "xs:list(A) \<Longrightarrow> \<forall>n\<in>nat. sublist(xs,n) = take(n,xs)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1240
apply (erule list.induct, simp)
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 46953
diff changeset
  1241
apply (clarify )
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1242
apply (erule natE)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1243
apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1244
done
71082
995fe5877d53 tuned declarations for more compact proof terms;
wenzelm
parents: 69587
diff changeset
  1245
declare sublist_upt_eq_take [simp]
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1246
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1247
lemma sublist_Int_eq:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1248
     "xs \<in> list(B) \<Longrightarrow> sublist(xs, A \<inter> nat) = sublist(xs, A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1249
apply (erule list.induct)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
  1250
apply (simp_all add: sublist_Cons)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1251
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1252
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
  1253
text\<open>Repetition of a List Element\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1254
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
  1255
consts   repeat :: "[i,i]\<Rightarrow>i"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1256
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1257
  "repeat(a,0) = []"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1258
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1259
  "repeat(a,succ(n)) = Cons(a,repeat(a,n))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1260
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1261
lemma length_repeat: "n \<in> nat \<Longrightarrow> length(repeat(a,n)) = n"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1262
by (induct_tac n, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1263
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1264
lemma repeat_succ_app: "n \<in> nat \<Longrightarrow> repeat(a,succ(n)) = repeat(a,n) @ [a]"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1265
apply (induct_tac n)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1266
apply (simp_all del: app_Cons add: app_Cons [symmetric])
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1267
done
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1268
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71089
diff changeset
  1269
lemma repeat_type [TC]: "\<lbrakk>a \<in> A; n \<in> nat\<rbrakk> \<Longrightarrow> repeat(a,n) \<in> list(A)"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1270
by (induct_tac n, auto)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1271
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
  1272
end