src/HOL/Library/Tree_Real.thy
author wenzelm
Sun, 13 Dec 2020 23:11:41 +0100
changeset 72907 3883f536d84d
parent 72566 831f17da1aab
permissions -rw-r--r--
unused (see 29566b6810f7);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     2
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     3
theory Tree_Real
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     4
imports
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     5
  Complex_Main
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     6
  Tree
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     7
begin
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
     8
68484
59793df7f853 clarified document antiquotation @{theory};
wenzelm
parents: 67399
diff changeset
     9
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69117
diff changeset
    10
  This theory is separate from \<^theory>\<open>HOL-Library.Tree\<close> because the former is discrete and
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69117
diff changeset
    11
  builds on \<^theory>\<open>Main\<close> whereas this theory builds on \<^theory>\<open>Complex_Main\<close>.
68484
59793df7f853 clarified document antiquotation @{theory};
wenzelm
parents: 67399
diff changeset
    12
\<close>
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    13
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    14
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    15
lemma size1_height_log: "log 2 (size1 t) \<le> height t"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    16
by (simp add: log2_of_power_le size1_height)
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    17
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    18
lemma min_height_size1_log: "min_height t \<le> log 2 (size1 t)"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    19
by (simp add: le_log2_of_power min_height_size1)
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    20
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    21
lemma size1_log_if_complete: "complete t \<Longrightarrow> height t = log 2 (size1 t)"
70350
571ae57313a4 moved some theorems into HOL main corpus
haftmann
parents: 69593
diff changeset
    22
by (simp add: size1_if_complete)
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    23
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    24
lemma min_height_size1_log_if_incomplete:
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    25
  "\<not> complete t \<Longrightarrow> min_height t < log 2 (size1 t)"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    26
by (simp add: less_log2_of_power min_height_size1_if_incomplete)
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    27
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    28
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    29
lemma min_height_acomplete: assumes "acomplete t"
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    30
shows "min_height t = nat(floor(log 2 (size1 t)))"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    31
proof cases
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    32
  assume *: "complete t"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    33
  hence "size1 t = 2 ^ min_height t"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    34
    by (simp add: complete_iff_height size1_if_complete)
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    35
  from log2_of_power_eq[OF this] show ?thesis by linarith
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    36
next
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    37
  assume *: "\<not> complete t"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    38
  hence "height t = min_height t + 1"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    39
    using assms min_height_le_height[of t]
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    40
    by(auto simp: acomplete_def complete_iff_height)
69117
3d3e87835ae8 simplified proofs
nipkow
parents: 68998
diff changeset
    41
  hence "size1 t < 2 ^ (min_height t + 1)" by (metis * size1_height_if_incomplete)
3d3e87835ae8 simplified proofs
nipkow
parents: 68998
diff changeset
    42
  from floor_log_nat_eq_if[OF min_height_size1 this] show ?thesis by simp
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    43
qed
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    44
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    45
lemma height_acomplete: assumes "acomplete t"
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    46
shows "height t = nat(ceiling(log 2 (size1 t)))"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    47
proof cases
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    48
  assume *: "complete t"
69117
3d3e87835ae8 simplified proofs
nipkow
parents: 68998
diff changeset
    49
  hence "size1 t = 2 ^ height t" by (simp add: size1_if_complete)
3d3e87835ae8 simplified proofs
nipkow
parents: 68998
diff changeset
    50
  from log2_of_power_eq[OF this] show ?thesis by linarith
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    51
next
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    52
  assume *: "\<not> complete t"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    53
  hence **: "height t = min_height t + 1"
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    54
    using assms min_height_le_height[of t]
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    55
    by(auto simp add: acomplete_def complete_iff_height)
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    56
  hence "size1 t \<le> 2 ^ (min_height t + 1)" by (metis size1_height)
69117
3d3e87835ae8 simplified proofs
nipkow
parents: 68998
diff changeset
    57
  from log2_of_power_le[OF this size1_ge0] min_height_size1_log_if_incomplete[OF *] **
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    58
  show ?thesis by linarith
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    59
qed
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents:
diff changeset
    60
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    61
lemma acomplete_Node_if_wbal1:
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    62
assumes "acomplete l" "acomplete r" "size l = size r + 1"
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    63
shows "acomplete \<langle>l, x, r\<rangle>"
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    64
proof -
68998
818898556504 more traditional formulation
nipkow
parents: 68484
diff changeset
    65
  from assms(3) have [simp]: "size1 l = size1 r + 1" by(simp add: size1_size)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    66
  have "nat \<lceil>log 2 (1 + size1 r)\<rceil> \<ge> nat \<lceil>log 2 (size1 r)\<rceil>"
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    67
    by(rule nat_mono[OF ceiling_mono]) simp
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    68
  hence 1: "height(Node l x r) = nat \<lceil>log 2 (1 + size1 r)\<rceil> + 1"
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    69
    using height_acomplete[OF assms(1)] height_acomplete[OF assms(2)]
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    70
    by (simp del: nat_ceiling_le_eq add: max_def)
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    71
  have "nat \<lfloor>log 2 (1 + size1 r)\<rfloor> \<ge> nat \<lfloor>log 2 (size1 r)\<rfloor>"
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    72
    by(rule nat_mono[OF floor_mono]) simp
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    73
  hence 2: "min_height(Node l x r) = nat \<lfloor>log 2 (size1 r)\<rfloor> + 1"
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    74
    using min_height_acomplete[OF assms(1)] min_height_acomplete[OF assms(2)]
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    75
    by (simp)
68998
818898556504 more traditional formulation
nipkow
parents: 68484
diff changeset
    76
  have "size1 r \<ge> 1" by(simp add: size1_size)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    77
  then obtain i where i: "2 ^ i \<le> size1 r" "size1 r < 2 ^ (i + 1)"
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    78
    using ex_power_ivl1[of 2 "size1 r"] by auto
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    79
  hence i1: "2 ^ i < size1 r + 1" "size1 r + 1 \<le> 2 ^ (i + 1)" by auto
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    80
  from 1 2 floor_log_nat_eq_if[OF i] ceiling_log_nat_eq_if[OF i1]
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    81
  show ?thesis by(simp add:acomplete_def)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    82
qed
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    83
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    84
lemma acomplete_sym: "acomplete \<langle>l, x, r\<rangle> \<Longrightarrow> acomplete \<langle>r, y, l\<rangle>"
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    85
by(auto simp: acomplete_def)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    86
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    87
lemma acomplete_Node_if_wbal2:
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    88
assumes "acomplete l" "acomplete r" "abs(int(size l) - int(size r)) \<le> 1"
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    89
shows "acomplete \<langle>l, x, r\<rangle>"
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    90
proof -
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    91
  have "size l = size r \<or> (size l = size r + 1 \<or> size r = size l + 1)" (is "?A \<or> ?B")
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    92
    using assms(3) by linarith
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    93
  thus ?thesis
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    94
  proof
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    95
    assume "?A"
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    96
    thus ?thesis using assms(1,2)
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    97
      apply(simp add: acomplete_def min_def max_def)
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
    98
      by (metis assms(1,2) acomplete_optimal le_antisym le_less)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
    99
  next
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   100
    assume "?B"
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   101
    thus ?thesis
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
   102
      by (meson assms(1,2) acomplete_sym acomplete_Node_if_wbal1)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   103
  qed
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   104
qed
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   105
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
   106
lemma acomplete_if_wbalanced: "wbalanced t \<Longrightarrow> acomplete t"
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   107
proof(induction t)
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
   108
  case Leaf show ?case by (simp add: acomplete_def)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   109
next
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   110
  case (Node l x r)
72566
831f17da1aab renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents: 70350
diff changeset
   111
  thus ?case by(simp add: acomplete_Node_if_wbal2)
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   112
qed
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   113
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66515
diff changeset
   114
end