src/HOL/Data_Structures/Sorted_Less.thy
author wenzelm
Sat, 31 Oct 2015 16:24:46 +0100
changeset 61530 aa1ece0bce62
parent 61203 a8a8eca85801
child 61640 44c9198f210c
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     2
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     3
section {* Lists Sorted wrt $<$ *}
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     4
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     5
theory Sorted_Less
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     6
imports Less_False
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     7
begin
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     8
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     9
hide_const sorted
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    10
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    11
text \<open>Is a list sorted without duplicates, i.e., wrt @{text"<"}?
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    12
Could go into theory List under a name like @{term sorted_less}.\<close>
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    13
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    14
fun sorted :: "'a::linorder list \<Rightarrow> bool" where
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    15
"sorted [] = True" |
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    16
"sorted [x] = True" |
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    17
"sorted (x#y#zs) = (x < y \<and> sorted(y#zs))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    18
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    19
lemma sorted_Cons_iff:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    20
  "sorted(x # xs) = (sorted xs \<and> (\<forall>y \<in> set xs. x < y))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    21
by(induction xs rule: sorted.induct) auto
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    22
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    23
lemma sorted_snoc_iff:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    24
  "sorted(xs @ [x]) = (sorted xs \<and> (\<forall>y \<in> set xs. y < x))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    25
by(induction xs rule: sorted.induct) auto
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    26
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    27
lemma sorted_cons: "sorted (x#xs) \<Longrightarrow> sorted xs"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    28
by(simp add: sorted_Cons_iff)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    29
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    30
lemma sorted_cons': "ASSUMPTION (sorted (x#xs)) \<Longrightarrow> sorted xs"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    31
by(rule ASSUMPTION_D [THEN sorted_cons])
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    32
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    33
lemma sorted_snoc: "sorted (xs @ [y]) \<Longrightarrow> sorted xs"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    34
by(simp add: sorted_snoc_iff)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    35
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    36
lemma sorted_snoc': "ASSUMPTION (sorted (xs @ [y])) \<Longrightarrow> sorted xs"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    37
by(rule ASSUMPTION_D [THEN sorted_snoc])
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    38
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    39
lemma sorted_mid_iff:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    40
  "sorted(xs @ y # ys) = (sorted(xs @ [y]) \<and> sorted(y # ys))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    41
by(induction xs rule: sorted.induct) auto
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    42
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    43
lemma sorted_mid_iff2:
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    44
  "sorted(x # xs @ y # ys) =
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    45
  (sorted(x # xs) \<and> x < y \<and> sorted(xs @ [y]) \<and> sorted(y # ys))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    46
by(induction xs rule: sorted.induct) auto
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    47
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    48
lemma sorted_mid_iff': "NO_MATCH [] ys \<Longrightarrow>
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    49
  sorted(xs @ y # ys) = (sorted(xs @ [y]) \<and> sorted(y # ys))"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    50
by(rule sorted_mid_iff)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    51
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    52
lemmas sorted_lems = sorted_mid_iff' sorted_mid_iff2 sorted_cons' sorted_snoc'
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    53
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    54
end