src/HOL/Induct/BinaryTree.thy
author paulson
Wed, 10 Dec 2003 15:59:34 +0100
changeset 14288 d149e3cbdb39
permissions -rw-r--r--
Moving some theorems from Real/RealArith0.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     1
header {* Isar-style Reasoning for Binary Tree Operations *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     2
theory BinaryTree = Main:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     3
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     4
text {* We prove correctness of operations on 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     5
 binary search tree implementing a set.
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     6
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     7
 This document is GPL.
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     8
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
     9
 Author: Viktor Kuncak, MIT CSAIL, November 2003 *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    10
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    11
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    12
section {* Tree Definition *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    13
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    14
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    15
datatype 'a Tree = Tip | T "'a Tree" 'a "'a Tree"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    16
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    17
consts
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    18
  setOf :: "'a Tree => 'a set" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    19
  -- {* set abstraction of a tree *} 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    20
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    21
"setOf Tip = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    22
"setOf (T t1 x t2) = (setOf t1) Un (setOf t2) Un {x}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    23
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    24
types
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    25
  -- {* we require index to have an irreflexive total order < *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    26
  -- {* apart from that, we do not rely on index being int *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    27
  index = int 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    28
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    29
types -- {* hash function type *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    30
  'a hash = "'a => index"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    31
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    32
constdefs
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    33
  eqs :: "'a hash => 'a => 'a set"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    34
  -- {* equivalence class of elements with the same hash code *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    35
  "eqs h x == {y. h y = h x}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    36
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    37
consts
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    38
  sortedTree :: "'a hash => 'a Tree => bool"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    39
  -- {* check if a tree is sorted *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    40
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    41
"sortedTree h Tip = True"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    42
"sortedTree h (T t1 x t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    43
  (sortedTree h t1 & 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    44
   (ALL l: setOf t1. h l < h x) &
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    45
   (ALL r: setOf t2. h x < h r) &
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    46
   sortedTree h t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    47
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    48
lemma sortLemmaL: 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    49
  "sortedTree h (T t1 x t2) ==> sortedTree h t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    50
lemma sortLemmaR: 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    51
  "sortedTree h (T t1 x t2) ==> sortedTree h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    52
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    53
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    54
section {* Tree Lookup *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    55
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    56
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    57
consts
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    58
  tlookup :: "'a hash => index => 'a Tree => 'a option"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    59
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    60
"tlookup h k Tip = None"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    61
"tlookup h k (T t1 x t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    62
 (if k < h x then tlookup h k t1
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    63
  else if h x < k then tlookup h k t2
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    64
  else Some x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    65
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    66
lemma tlookup_none: 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    67
"sortedTree h t & (tlookup h k t = None) -->
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    68
 (ALL x:setOf t. h x ~= k)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    69
apply (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    70
apply auto (* takes some time *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    71
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    72
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    73
lemma tlookup_some:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    74
"sortedTree h t & (tlookup h k t = Some x) -->
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    75
 x:setOf t & h x = k"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    76
apply (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    77
apply auto (* takes some time *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    78
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    79
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    80
constdefs
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    81
  sorted_distinct_pred :: "'a hash => 'a => 'a => 'a Tree => bool"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    82
  -- {* No two elements have the same hash code *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    83
  "sorted_distinct_pred h a b t == sortedTree h t & 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    84
      a:setOf t & b:setOf t & h a = h b --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    85
      a = b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    86
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    87
declare sorted_distinct_pred_def [simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    88
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    89
-- {* for case analysis on three cases *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    90
lemma cases3: "[| C1 ==> G; C2 ==> G; C3 ==> G;
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    91
                  C1 | C2 | C3 |] ==> G"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    92
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    93
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    94
text {* @{term sorted_distinct_pred} holds for out trees: *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    95
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    96
lemma sorted_distinct: "sorted_distinct_pred h a b t" (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    97
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    98
  show "?P Tip" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
    99
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   100
  fix t2 :: "'a Tree" assume h2: "?P t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   101
  fix x :: 'a
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   102
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   103
  proof (unfold sorted_distinct_pred_def, safe)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   104
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   105
    assume adef: "a : setOf (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   106
    assume bdef: "b : setOf (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   107
    assume hahb: "h a = h b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   108
    from s have s1: "sortedTree h t1" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   109
    from s have s2: "sortedTree h t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   110
    show "a = b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   111
    -- {* We consider 9 cases for the position of a and b are in the tree *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   112
    proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   113
    -- {* three cases for a *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   114
    from adef have "a : setOf t1 | a = x | a : setOf t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   115
    moreover { assume adef1: "a : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   116
      have ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   117
      proof - 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   118
      -- {* three cases for b *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   119
      from bdef have "b : setOf t1 | b = x | b : setOf t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   120
      moreover { assume bdef1: "b : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   121
        from s1 adef1 bdef1 hahb h1 have ?thesis by simp }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   122
      moreover { assume bdef1: "b = x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   123
        from adef1 bdef1 s have "h a < h b" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   124
        from this hahb have ?thesis by simp }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   125
      moreover { assume bdef1: "b : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   126
        from adef1 s have o1: "h a < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   127
        from bdef1 s have o2: "h x < h b" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   128
        from o1 o2 have "h a < h b" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   129
        from this hahb have ?thesis by simp } -- {* case impossible *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   130
      ultimately show ?thesis by blast
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   131
      qed 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   132
    } 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   133
    moreover { assume adef1: "a = x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   134
      have ?thesis 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   135
      proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   136
      -- {* three cases for b *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   137
      from bdef have "b : setOf t1 | b = x | b : setOf t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   138
      moreover { assume bdef1: "b : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   139
        from this s have "h b < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   140
        from this adef1 have "h b < h a" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   141
        from hahb this have ?thesis by simp } -- {* case impossible *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   142
      moreover { assume bdef1: "b = x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   143
        from adef1 bdef1 have ?thesis by simp }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   144
      moreover { assume bdef1: "b : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   145
        from this s have "h x < h b" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   146
        from this adef1 have "h a < h b" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   147
        from hahb this have ?thesis by simp } -- {* case impossible *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   148
      ultimately show ?thesis by blast
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   149
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   150
    }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   151
    moreover { assume adef1: "a : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   152
      have ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   153
      proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   154
      -- {* three cases for b *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   155
      from bdef have "b : setOf t1 | b = x | b : setOf t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   156
      moreover { assume bdef1: "b : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   157
        from bdef1 s have o1: "h b < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   158
        from adef1 s have o2: "h x < h a" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   159
        from o1 o2 have "h b < h a" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   160
        from this hahb have ?thesis by simp } -- {* case impossible *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   161
      moreover { assume bdef1: "b = x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   162
        from adef1 bdef1 s have "h b < h a" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   163
        from this hahb have ?thesis by simp } -- {* case impossible *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   164
      moreover { assume bdef1: "b : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   165
        from s2 adef1 bdef1 hahb h2 have ?thesis by simp }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   166
      ultimately show ?thesis by blast
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   167
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   168
    }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   169
    ultimately show ?thesis by blast
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   170
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   171
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   172
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   173
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   174
lemma tlookup_finds: -- {* if a node is in the tree, lookup finds it *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   175
"sortedTree h t & y:setOf t --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   176
 tlookup h (h y) t = Some y"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   177
proof safe
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   178
  assume s: "sortedTree h t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   179
  assume yint: "y : setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   180
  show "tlookup h (h y) t = Some y"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   181
  proof (cases "tlookup h (h y) t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   182
  case None note res = this    
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   183
    from s res have "sortedTree h t & (tlookup h (h y) t = None)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   184
    from this have o1: "ALL x:setOf t. h x ~= h y" by (simp add: tlookup_none)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   185
    from o1 yint have "h y ~= h y" by fastsimp (* auto does not work *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   186
    from this show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   187
  next case (Some z) note res = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   188
    have ls: "sortedTree h t & (tlookup h (h y) t = Some z) -->
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   189
              z:setOf t & h z = h y" by (simp add: tlookup_some)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   190
    have sd: "sorted_distinct_pred h y z t" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   191
    by (insert sorted_distinct [of h y z t], simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   192
       (* for some reason simplifier would never guess this substitution *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   193
    from s res ls have o1: "z:setOf t & h z = h y" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   194
    from s yint o1 sd have "y = z" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   195
    from this res show "tlookup h (h y) t = Some y" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   196
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   197
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   198
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   199
subsection {* Tree membership as a special case of lookup *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   200
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   201
constdefs
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   202
  memb :: "'a hash => 'a => 'a Tree => bool"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   203
  "memb h x t == 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   204
   (case (tlookup h (h x) t) of
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   205
      None => False
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   206
    | Some z => (x=z))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   207
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   208
lemma assumes s: "sortedTree h t" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   209
      shows memb_spec: "memb h x t = (x : setOf t)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   210
proof (cases "tlookup h (h x) t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   211
case None note tNone = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   212
  from tNone have res: "memb h x t = False" by (simp add: memb_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   213
  from s tNone tlookup_none have o1: "ALL y:setOf t. h y ~= h x" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   214
  have notIn: "x ~: setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   215
  proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   216
    assume h: "x : setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   217
    from h o1 have "h x ~= h x" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   218
    from this show False by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   219
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   220
  from res notIn show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   221
next case (Some z) note tSome = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   222
  from s tSome tlookup_some have zin: "z : setOf t" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   223
  show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   224
  proof (cases "x=z")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   225
  case True note xez = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   226
    from tSome xez have res: "memb h x t" by (simp add: memb_def)  
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   227
    from res zin xez show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   228
  next case False note xnez = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   229
    from tSome xnez have res: "~ memb h x t" by (simp add: memb_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   230
    have "x ~: setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   231
    proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   232
      assume xin: "x : setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   233
      from s tSome tlookup_some have hzhx: "h x = h z" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   234
      have o1: "sorted_distinct_pred h x z t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   235
      by (insert sorted_distinct [of h x z t], simp)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   236
      from s xin zin hzhx o1 have "x = z" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   237
      from this xnez show False by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   238
    qed  
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   239
    from this res show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   240
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   241
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   242
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   243
declare sorted_distinct_pred_def [simp del]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   244
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   245
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   246
section {* Insertion into a Tree *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   247
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   248
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   249
consts
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   250
  binsert :: "'a hash => 'a => 'a Tree => 'a Tree"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   251
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   252
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   253
"binsert h e Tip = (T Tip e Tip)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   254
"binsert h e (T t1 x t2) = (if h e < h x then
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   255
                             (T (binsert h e t1) x t2)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   256
                            else
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   257
                             (if h x < h e then
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   258
                               (T t1 x (binsert h e t2))
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   259
                              else (T t1 e t2)))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   260
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   261
text {* A technique for proving disjointness of sets. *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   262
lemma disjCond: "[| !! x. [| x:A; x:B |] ==> False |] ==> A Int B = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   263
by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   264
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   265
text {* The following is a proof that insertion correctly implements
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   266
        the set interface.
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   267
        Compared to @{text BinaryTree_TacticStyle}, the claim is more
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   268
        difficult, and this time we need to assume as a hypothesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   269
        that the tree is sorted. *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   270
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   271
lemma binsert_set: "sortedTree h t -->
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   272
                    setOf (binsert h e t) = (setOf t) - (eqs h e) Un {e}" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   273
      (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   274
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   275
  -- {* base case *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   276
  show "?P Tip" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   277
  -- {* inductition step *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   278
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   279
  fix t2 :: "'a Tree" assume h2: "?P t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   280
  fix x :: 'a
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   281
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   282
  proof 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   283
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   284
    from s have s1: "sortedTree h t1" by (rule sortLemmaL)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   285
    from s1 and h1 have c1: "setOf (binsert h e t1) = setOf t1 - eqs h e Un {e}" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   286
    from s have s2: "sortedTree h t2" by (rule sortLemmaR)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   287
    from s2 and h2 have c2: "setOf (binsert h e t2) = setOf t2 - eqs h e Un {e}" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   288
    show "setOf (binsert h e (T t1 x t2)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   289
          setOf (T t1 x t2) - eqs h e Un {e}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   290
    proof (cases "h e < h x")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   291
      case True note eLess = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   292
        from eLess have res: "binsert h e (T t1 x t2) = (T (binsert h e t1) x t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   293
        show "setOf (binsert h e (T t1 x t2)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   294
              setOf (T t1 x t2) - eqs h e Un {e}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   295
        proof (simp add: res eLess c1)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   296
          show "insert x (insert e (setOf t1 - eqs h e Un setOf t2)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   297
                insert e (insert x (setOf t1 Un setOf t2) - eqs h e)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   298
          proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   299
            have eqsLessX: "ALL el: eqs h e. h el < h x" by (simp add: eqs_def eLess)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   300
            from this have eqsDisjX: "ALL el: eqs h e. h el ~= h x" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   301
            from s have xLessT2: "ALL r: setOf t2. h x < h r" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   302
            have eqsLessT2: "ALL el: eqs h e. ALL r: setOf t2. h el < h r"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   303
            proof safe
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   304
              fix el assume hel: "el : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   305
              from hel eqs_def have o1: "h el = h e" by fastsimp (* auto fails here! *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   306
              fix r assume hr: "r : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   307
              from xLessT2 hr o1 eLess show "h el < h r" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   308
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   309
            from eqsLessT2 have eqsDisjT2: "ALL el: eqs h e. ALL r: setOf t2. h el ~= h r"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   310
            by fastsimp (* auto fails here *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   311
            from eqsDisjX eqsDisjT2 show ?thesis by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   312
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   313
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   314
      next case False note eNotLess = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   315
      show "setOf (binsert h e (T t1 x t2)) = setOf (T t1 x t2) - eqs h e Un {e}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   316
      proof (cases "h x < h e")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   317
        case True note xLess = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   318
        from xLess have res: "binsert h e (T t1 x t2) = (T t1 x (binsert h e t2))" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   319
        show "setOf (binsert h e (T t1 x t2)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   320
              setOf (T t1 x t2) - eqs h e Un {e}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   321
        proof (simp add: res xLess eNotLess c2)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   322
          show "insert x (insert e (setOf t1 Un (setOf t2 - eqs h e))) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   323
                insert e (insert x (setOf t1 Un setOf t2) - eqs h e)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   324
          proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   325
            have XLessEqs: "ALL el: eqs h e. h x < h el" by (simp add: eqs_def xLess)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   326
            from this have eqsDisjX: "ALL el: eqs h e. h el ~= h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   327
            from s have t1LessX: "ALL l: setOf t1. h l < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   328
            have T1lessEqs: "ALL el: eqs h e. ALL l: setOf t1. h l < h el"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   329
            proof safe
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   330
              fix el assume hel: "el : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   331
              fix l assume hl: "l : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   332
              from hel eqs_def have o1: "h el = h e" by fastsimp (* auto fails here! *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   333
              from t1LessX hl o1 xLess show "h l < h el" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   334
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   335
            from T1lessEqs have T1disjEqs: "ALL el: eqs h e. ALL l: setOf t1. h el ~= h l"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   336
            by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   337
            from eqsDisjX T1lessEqs show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   338
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   339
        qed      
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   340
      next case False note xNotLess = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   341
        from xNotLess eNotLess have xeqe: "h x = h e" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   342
        from xeqe have res: "binsert h e (T t1 x t2) = (T t1 e t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   343
        show "setOf (binsert h e (T t1 x t2)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   344
              setOf (T t1 x t2) - eqs h e Un {e}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   345
        proof (simp add: res eNotLess xeqe)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   346
          show "insert e (setOf t1 Un setOf t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   347
                insert e (insert x (setOf t1 Un setOf t2) - eqs h e)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   348
          proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   349
            have "insert x (setOf t1 Un setOf t2) - eqs h e = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   350
                  setOf t1 Un setOf t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   351
            proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   352
              have (* o1: *) "x : eqs h e" by (simp add: eqs_def xeqe)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   353
              moreover have (* o2: *) "(setOf t1) Int (eqs h e) = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   354
              proof (rule disjCond)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   355
                fix w
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   356
                assume whSet: "w : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   357
                assume whEq: "w : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   358
                from whSet s have o1: "h w < h x" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   359
                from whEq eqs_def have o2: "h w = h e" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   360
                from o2 xeqe have o3: "~ h w < h x" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   361
                from o1 o3 show False by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   362
              qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   363
              moreover have (* o3: *) "(setOf t2) Int (eqs h e) = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   364
              proof (rule disjCond)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   365
                fix w
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   366
                assume whSet: "w : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   367
                assume whEq: "w : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   368
                from whSet s have o1: "h x < h w" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   369
                from whEq eqs_def have o2: "h w = h e" by fastsimp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   370
                from o2 xeqe have o3: "~ h x < h w" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   371
                from o1 o3 show False by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   372
              qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   373
              ultimately show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   374
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   375
            from this show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   376
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   377
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   378
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   379
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   380
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   381
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   382
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   383
text {* Using the correctness of set implementation,
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   384
        preserving sortedness is still simple. *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   385
lemma binsert_sorted: "sortedTree h t --> sortedTree h (binsert h x t)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   386
by (induct t) (auto simp add: binsert_set)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   387
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   388
text {* We summarize the specification of binsert as follows. *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   389
corollary binsert_spec: "sortedTree h t -->
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   390
                     sortedTree h (binsert h x t) &
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   391
                     setOf (binsert h e t) = (setOf t) - (eqs h e) Un {e}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   392
by (simp add: binsert_set binsert_sorted)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   393
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   394
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   395
section {* Removing an element from a tree *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   396
(*============================================================*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   397
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   398
text {* These proofs are influenced by those in @{text BinaryTree_Tactic} *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   399
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   400
consts 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   401
  rm :: "'a hash => 'a Tree => 'a"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   402
  -- {* rightmost element of a tree *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   403
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   404
"rm h (T t1 x t2) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   405
  (if t2=Tip then x else rm h t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   406
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   407
consts 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   408
  wrm :: "'a hash => 'a Tree => 'a Tree"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   409
  -- {* tree without the rightmost element *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   410
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   411
"wrm h (T t1 x t2) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   412
  (if t2=Tip then t1 else (T t1 x (wrm h t2)))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   413
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   414
consts 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   415
  wrmrm :: "'a hash => 'a Tree => 'a Tree * 'a"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   416
  -- {* computing rightmost and removal in one pass *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   417
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   418
"wrmrm h (T t1 x t2) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   419
  (if t2=Tip then (t1,x)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   420
   else (T t1 x (fst (wrmrm h t2)),
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   421
         snd (wrmrm h t2)))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   422
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   423
consts
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   424
  remove :: "'a hash => 'a => 'a Tree => 'a Tree"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   425
   -- {* removal of an element from the tree *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   426
primrec
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   427
"remove h e Tip = Tip"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   428
"remove h e (T t1 x t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   429
  (if h e < h x then (T (remove h e t1) x t2)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   430
   else if h x < h e then (T t1 x (remove h e t2))
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   431
   else (if t1=Tip then t2
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   432
         else let (t1p,r) = wrmrm h t1
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   433
              in (T t1p r t2)))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   434
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   435
theorem wrmrm_decomp: "t ~= Tip --> wrmrm h t = (wrm h t, rm h t)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   436
apply (induct_tac t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   437
apply simp_all
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   438
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   439
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   440
lemma rm_set: "t ~= Tip & sortedTree h t --> rm h t : setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   441
apply (induct_tac t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   442
apply simp_all
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   443
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   444
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   445
lemma wrm_set: "t ~= Tip & sortedTree h t --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   446
                setOf (wrm h t) = setOf t - {rm h t}" (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   447
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   448
  show "?P Tip" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   449
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   450
  fix t2 :: "'a Tree" assume h2: "?P t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   451
  fix x :: 'a
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   452
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   453
  proof (rule impI, erule conjE)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   454
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   455
    show "setOf (wrm h (T t1 x t2)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   456
          setOf (T t1 x t2) - {rm h (T t1 x t2)}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   457
    proof (cases "t2 = Tip")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   458
    case True note t2tip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   459
      from t2tip have rm_res: "rm h (T t1 x t2) = x" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   460
      from t2tip have wrm_res: "wrm h (T t1 x t2) = t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   461
      from s have "x ~: setOf t1" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   462
      from this rm_res wrm_res t2tip show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   463
    next case False note t2nTip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   464
      from t2nTip have rm_res: "rm h (T t1 x t2) = rm h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   465
      from t2nTip have wrm_res: "wrm h (T t1 x t2) = T t1 x (wrm h t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   466
      from s have s2: "sortedTree h t2" by simp    
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   467
      from h2 t2nTip s2 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   468
      have o1: "setOf (wrm h t2) = setOf t2 - {rm h t2}" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   469
      show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   470
      proof (simp add: rm_res wrm_res t2nTip h2 o1)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   471
        show "insert x (setOf t1 Un (setOf t2 - {rm h t2})) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   472
              insert x (setOf t1 Un setOf t2) - {rm h t2}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   473
        proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   474
          from s rm_set t2nTip have xOk: "h x < h (rm h t2)" by auto 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   475
          have t1Ok: "ALL l:setOf t1. h l < h (rm h t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   476
          proof safe
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   477
            fix l :: 'a  assume ldef: "l : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   478
            from ldef s have lx: "h l < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   479
            from lx xOk show "h l < h (rm h t2)" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   480
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   481
          from xOk t1Ok show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   482
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   483
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   484
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   485
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   486
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   487
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   488
lemma wrm_set1: "t ~= Tip & sortedTree h t --> setOf (wrm h t) <= setOf t"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   489
by (auto simp add: wrm_set)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   490
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   491
lemma wrm_sort: "t ~= Tip & sortedTree h t --> sortedTree h (wrm h t)" (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   492
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   493
  show "?P Tip" by simp  
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   494
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   495
  fix t2 :: "'a Tree" assume h2: "?P t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   496
  fix x :: 'a
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   497
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   498
  proof safe
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   499
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   500
    show "sortedTree h (wrm h (T t1 x t2))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   501
    proof (cases "t2 = Tip")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   502
    case True note t2tip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   503
      from t2tip have res: "wrm h (T t1 x t2) = t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   504
      from res s show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   505
    next case False note t2nTip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   506
      from t2nTip have res: "wrm h (T t1 x t2) = T t1 x (wrm h t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   507
      from s have s1: "sortedTree h t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   508
      from s have s2: "sortedTree h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   509
      from s2 h2 t2nTip have o1: "sortedTree h (wrm h t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   510
      from s2 t2nTip wrm_set1 have o2: "setOf (wrm h t2) <= setOf t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   511
      from s o2 have o3: "ALL r: setOf (wrm h t2). h x < h r" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   512
      from s1 o1 o3 res s show "sortedTree h (wrm h (T t1 x t2))" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   513
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   514
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   515
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   516
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   517
lemma wrm_less_rm: 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   518
  "t ~= Tip & sortedTree h t --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   519
   (ALL l:setOf (wrm h t). h l < h (rm h t))" (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   520
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   521
  show "?P Tip" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   522
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   523
  fix t2 :: "'a Tree" assume h2: "?P t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   524
  fix x :: 'a   
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   525
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   526
  proof safe 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   527
    fix l :: "'a" assume ldef: "l : setOf (wrm h (T t1 x t2))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   528
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   529
    from s have s1: "sortedTree h t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   530
    from s have s2: "sortedTree h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   531
    show "h l < h (rm h (T t1 x t2))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   532
    proof (cases "t2 = Tip")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   533
    case True note t2tip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   534
      from t2tip have rm_res: "rm h (T t1 x t2) = x" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   535
      from t2tip have wrm_res: "wrm h (T t1 x t2) = t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   536
      from ldef wrm_res have o1: "l : setOf t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   537
      from rm_res o1 s show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   538
    next case False note t2nTip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   539
      from t2nTip have rm_res: "rm h (T t1 x t2) = rm h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   540
      from t2nTip have wrm_res: "wrm h (T t1 x t2) = T t1 x (wrm h t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   541
      from ldef wrm_res 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   542
      have l_scope: "l : {x} Un setOf t1 Un setOf (wrm h t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   543
      have hLess: "h l < h (rm h t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   544
      proof (cases "l = x")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   545
      case True note lx = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   546
        from s t2nTip rm_set s2 have o1: "h x < h (rm h t2)" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   547
        from lx o1 show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   548
      next case False note lnx = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   549
        show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   550
        proof (cases "l : setOf t1")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   551
        case True note l_in_t1 = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   552
          from s t2nTip rm_set s2 have o1: "h x < h (rm h t2)" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   553
          from l_in_t1 s have o2: "h l < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   554
          from o1 o2 show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   555
        next case False note l_notin_t1 = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   556
          from l_scope lnx l_notin_t1 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   557
          have l_in_res: "l : setOf (wrm h t2)" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   558
          from l_in_res h2 t2nTip s2 show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   559
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   560
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   561
      from rm_res hLess show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   562
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   563
  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   564
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   565
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   566
lemma remove_set: "sortedTree h t --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   567
  setOf (remove h e t) = setOf t - eqs h e" (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   568
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   569
  show "?P Tip" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   570
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   571
  fix t2 :: "'a Tree" assume h2: "?P t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   572
  fix x :: 'a
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   573
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   574
  proof 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   575
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   576
    show "setOf (remove h e (T t1 x t2)) = setOf (T t1 x t2) - eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   577
    proof (cases "h e < h x")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   578
    case True note elx = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   579
      from elx have res: "remove h e (T t1 x t2) = T (remove h e t1) x t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   580
      by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   581
      from s have s1: "sortedTree h t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   582
      from s1 h1 have o1: "setOf (remove h e t1) = setOf t1 - eqs h e" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   583
      show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   584
      proof (simp add: o1 elx)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   585
        show "insert x (setOf t1 - eqs h e Un setOf t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   586
              insert x (setOf t1 Un setOf t2) - eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   587
        proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   588
          have xOk: "x ~: eqs h e" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   589
          proof 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   590
            assume h: "x : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   591
            from h have o1: "~ (h e < h x)" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   592
            from elx o1 show "False" by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   593
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   594
          have t2Ok: "(setOf t2) Int (eqs h e) = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   595
          proof (rule disjCond)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   596
            fix y :: 'a 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   597
            assume y_in_t2: "y : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   598
            assume y_in_eq: "y : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   599
            from y_in_t2 s have xly: "h x < h y" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   600
            from y_in_eq have eey: "h y = h e" by (simp add: eqs_def) (* must "add:" not "from" *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   601
            from xly eey have nelx: "~ (h e < h x)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   602
            from nelx elx show False by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   603
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   604
          from xOk t2Ok show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   605
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   606
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   607
    next case False note nelx = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   608
      show ?thesis 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   609
      proof (cases "h x < h e")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   610
      case True note xle = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   611
        from xle have res: "remove h e (T t1 x t2) = T t1 x (remove h e t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   612
        from s have s2: "sortedTree h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   613
        from s2 h2 have o1: "setOf (remove h e t2) = setOf t2 - eqs h e" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   614
        show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   615
        proof (simp add: o1 xle nelx)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   616
          show "insert x (setOf t1 Un (setOf t2 - eqs h e)) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   617
                insert x (setOf t1 Un setOf t2) - eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   618
          proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   619
            have xOk: "x ~: eqs h e" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   620
            proof 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   621
              assume h: "x : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   622
              from h have o1: "~ (h x < h e)" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   623
              from xle o1 show "False" by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   624
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   625
            have t1Ok: "(setOf t1) Int (eqs h e) = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   626
            proof (rule disjCond)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   627
              fix y :: 'a 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   628
              assume y_in_t1: "y : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   629
              assume y_in_eq: "y : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   630
              from y_in_t1 s have ylx: "h y < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   631
              from y_in_eq have eey: "h y = h e" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   632
              from ylx eey have nxle: "~ (h x < h e)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   633
              from nxle xle show False by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   634
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   635
            from xOk t1Ok show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   636
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   637
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   638
      next case False note nxle = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   639
        from nelx nxle have ex: "h e = h x" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   640
        have t2Ok: "(setOf t2) Int (eqs h e) = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   641
        proof (rule disjCond)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   642
          fix y :: 'a 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   643
          assume y_in_t2: "y : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   644
          assume y_in_eq: "y : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   645
          from y_in_t2 s have xly: "h x < h y" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   646
          from y_in_eq have eey: "h y = h e" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   647
          from y_in_eq ex eey have nxly: "~ (h x < h y)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   648
          from nxly xly show False by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   649
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   650
        show ?thesis 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   651
        proof (cases "t1 = Tip")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   652
        case True note t1tip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   653
          from ex t1tip have res: "remove h e (T t1 x t2) = t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   654
          show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   655
          proof (simp add: res t1tip ex)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   656
            show "setOf t2 = insert x (setOf t2) - eqs h e"              
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   657
            proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   658
              from ex have x_in_eqs: "x : eqs h e" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   659
              from x_in_eqs t2Ok show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   660
           qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   661
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   662
        next case False note t1nTip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   663
          from nelx nxle ex t1nTip
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   664
          have res: "remove h e (T t1 x t2) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   665
                     T (wrm h t1) (rm h t1) t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   666
          by (simp add: Let_def wrmrm_decomp)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   667
          from res show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   668
          proof simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   669
            from s have s1: "sortedTree h t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   670
            show "insert (rm h t1) (setOf (wrm h t1) Un setOf t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   671
                  insert x (setOf t1 Un setOf t2) - eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   672
            proof (simp add: t1nTip s1 rm_set wrm_set)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   673
              show "insert (rm h t1) (setOf t1 - {rm h t1} Un setOf t2) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   674
                    insert x (setOf t1 Un setOf t2) - eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   675
              proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   676
                from t1nTip s1 rm_set
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   677
                have o1: "insert (rm h t1) (setOf t1 - {rm h t1} Un setOf t2) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   678
                          setOf t1 Un setOf t2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   679
                have o2: "insert x (setOf t1 Un setOf t2) - eqs h e =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   680
                          setOf t1 Un setOf t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   681
                proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   682
                  from ex have xOk: "x : eqs h e" by (simp add: eqs_def)                  
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   683
		  have t1Ok: "(setOf t1) Int (eqs h e) = {}"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   684
                  proof (rule disjCond)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   685
                    fix y :: 'a 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   686
                    assume y_in_t1: "y : setOf t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   687
                    assume y_in_eq: "y : eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   688
                    from y_in_t1 s ex have o1: "h y < h e" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   689
                    from y_in_eq have o2: "~ (h y < h e)" by (simp add: eqs_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   690
                    from o1 o2 show False by contradiction
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   691
                  qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   692
                  from xOk t1Ok t2Ok show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   693
                qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   694
                from o1 o2 show ?thesis by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   695
              qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   696
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   697
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   698
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   699
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   700
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   701
  qed  
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   702
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   703
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   704
lemma remove_sort: "sortedTree h t --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   705
                    sortedTree h (remove h e t)" (is "?P t")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   706
proof (induct t)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   707
  show "?P Tip" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   708
  fix t1 :: "'a Tree" assume h1: "?P t1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   709
  fix t2 :: "'a Tree" assume h2: "?P t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   710
  fix x :: 'a
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   711
  show "?P (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   712
  proof 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   713
    assume s: "sortedTree h (T t1 x t2)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   714
    from s have s1: "sortedTree h t1" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   715
    from s have s2: "sortedTree h t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   716
    from h1 s1 have sr1: "sortedTree h (remove h e t1)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   717
    from h2 s2 have sr2: "sortedTree h (remove h e t2)" by simp   
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   718
    show "sortedTree h (remove h e (T t1 x t2))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   719
    proof (cases "h e < h x")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   720
    case True note elx = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   721
      from elx have res: "remove h e (T t1 x t2) = T (remove h e t1) x t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   722
      by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   723
      show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   724
      proof (simp add: s sr1 s2 elx res)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   725
        let ?C1 = "ALL l:setOf (remove h e t1). h l < h x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   726
        let ?C2 = "ALL r:setOf t2. h x < h r"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   727
        have o1: "?C1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   728
        proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   729
          from s1 have "setOf (remove h e t1) = setOf t1 - eqs h e" by (simp add: remove_set)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   730
          from s this show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   731
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   732
        from o1 s show "?C1 & ?C2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   733
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   734
    next case False note nelx = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   735
      show ?thesis 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   736
      proof (cases "h x < h e")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   737
      case True note xle = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   738
        from xle have res: "remove h e (T t1 x t2) = T t1 x (remove h e t2)" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   739
        show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   740
        proof (simp add: s s1 sr2 xle nelx res)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   741
          let ?C1 = "ALL l:setOf t1. h l < h x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   742
          let ?C2 = "ALL r:setOf (remove h e t2). h x < h r"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   743
          have o2: "?C2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   744
          proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   745
            from s2 have "setOf (remove h e t2) = setOf t2 - eqs h e" by (simp add: remove_set)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   746
            from s this show ?thesis by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   747
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   748
          from o2 s show "?C1 & ?C2" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   749
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   750
      next case False note nxle = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   751
        from nelx nxle have ex: "h e = h x" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   752
        show ?thesis 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   753
        proof (cases "t1 = Tip")
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   754
        case True note t1tip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   755
          from ex t1tip have res: "remove h e (T t1 x t2) = t2" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   756
          show ?thesis by (simp add: res t1tip ex s2)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   757
        next case False note t1nTip = this
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   758
          from nelx nxle ex t1nTip
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   759
          have res: "remove h e (T t1 x t2) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   760
                     T (wrm h t1) (rm h t1) t2" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   761
          by (simp add: Let_def wrmrm_decomp)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   762
          from res show ?thesis
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   763
          proof simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   764
            let ?C1 = "sortedTree h (wrm h t1)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   765
            let ?C2 = "ALL l:setOf (wrm h t1). h l < h (rm h t1)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   766
            let ?C3 = "ALL r:setOf t2. h (rm h t1) < h r"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   767
            let ?C4 = "sortedTree h t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   768
            from s1 t1nTip have o1: ?C1 by (simp add: wrm_sort)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   769
            from s1 t1nTip have o2: ?C2 by (simp add: wrm_less_rm)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   770
            have o3: ?C3
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   771
            proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   772
              fix r :: 'a 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   773
              assume rt2: "r : setOf t2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   774
              from s rm_set s1 t1nTip have o1: "h (rm h t1) < h x" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   775
              from rt2 s have o2: "h x < h r" by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   776
              from o1 o2 show "h (rm h t1) < h r" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   777
            qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   778
            from o1 o2 o3 s2 show "?C1 & ?C2 & ?C3 & ?C4" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   779
          qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   780
        qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   781
      qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   782
    qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   783
  qed  
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   784
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   785
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   786
text {* We summarize the specification of remove as follows. *}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   787
corollary remove_spec: "sortedTree h t --> 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   788
     sortedTree h (remove h e t) &
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   789
     setOf (remove h e t) = setOf t - eqs h e"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   790
by (simp add: remove_sort remove_set)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   791
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   792
generate_code ("BinaryTree_Code.ML") 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   793
  test = "tlookup id 4 (remove id 3 (binsert id 4 (binsert id 3 Tip)))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   794
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents:
diff changeset
   795
end