src/HOL/SizeChange/Misc_Tools.thy
author ballarin
Thu, 11 Dec 2008 18:30:26 +0100
changeset 29223 e09c53289830
parent 25314 5eaf3e8b50a4
permissions -rw-r--r--
Conversion of HOL-Main and ZF to new locales.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Library/SCT_Misc.thy
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     2
    ID:         $Id$
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     4
*)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     5
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     6
header {* Miscellaneous Tools for Size-Change Termination *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     7
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     8
theory Misc_Tools
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     9
imports Main
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    10
begin
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    11
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    12
subsection {* Searching in lists *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    13
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    14
fun index_of :: "'a list \<Rightarrow> 'a \<Rightarrow> nat"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    15
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    16
  "index_of [] c = 0"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    17
| "index_of (x#xs) c = (if x = c then 0 else Suc (index_of xs c))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    18
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    19
lemma index_of_member: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    20
  "(x \<in> set l) \<Longrightarrow> (l ! index_of l x = x)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    21
  by (induct l) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    22
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    23
lemma index_of_length:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    24
  "(x \<in> set l) = (index_of l x < length l)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    25
  by (induct l) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    26
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    27
subsection {* Some reasoning tools *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    28
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    29
lemma three_cases:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    30
  assumes "a1 \<Longrightarrow> thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    31
  assumes "a2 \<Longrightarrow> thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    32
  assumes "a3 \<Longrightarrow> thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    33
  assumes "\<And>R. \<lbrakk>a1 \<Longrightarrow> R; a2 \<Longrightarrow> R; a3 \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    34
  shows "thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    35
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    36
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    37
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    38
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    39
subsection {* Sequences *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    40
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    41
types
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    42
  'a sequence = "nat \<Rightarrow> 'a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    43
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    44
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    45
subsubsection {* Increasing sequences *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    46
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    47
definition
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    48
  increasing :: "(nat \<Rightarrow> nat) \<Rightarrow> bool" where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    49
  "increasing s = (\<forall>i j. i < j \<longrightarrow> s i < s j)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    50
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    51
lemma increasing_strict:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    52
  assumes "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    53
  assumes "i < j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    54
  shows "s i < s j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    55
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    56
  unfolding increasing_def by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    57
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    58
lemma increasing_weak:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    59
  assumes "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    60
  assumes "i \<le> j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    61
  shows "s i \<le> s j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    62
  using assms increasing_strict[of s i j]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    63
  by (cases "i < j") auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    64
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    65
lemma increasing_inc:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    66
  assumes "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    67
  shows "n \<le> s n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    68
proof (induct n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    69
  case 0 then show ?case by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    70
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    71
  case (Suc n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    72
  with increasing_strict [OF `increasing s`, of n "Suc n"]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    73
  show ?case by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    74
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    75
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    76
lemma increasing_bij:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    77
  assumes [simp]: "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    78
  shows "(s i < s j) = (i < j)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    79
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    80
  assume "s i < s j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    81
  show "i < j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    82
  proof (rule classical)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    83
    assume "\<not> ?thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    84
    hence "j \<le> i" by arith
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    85
    with increasing_weak have "s j \<le> s i" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    86
    with `s i < s j` show ?thesis by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    87
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    88
qed (simp add:increasing_strict)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    89
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    90
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    91
subsubsection {* Sections induced by an increasing sequence *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    92
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    93
abbreviation
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    94
  "section s i == {s i ..< s (Suc i)}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    95
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    96
definition
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    97
  "section_of s n = (LEAST i. n < s (Suc i))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    98
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    99
lemma section_help:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   100
  assumes "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   101
  shows "\<exists>i. n < s (Suc i)" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   102
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   103
  have "n \<le> s n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   104
    using `increasing s` by (rule increasing_inc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   105
  also have "\<dots> < s (Suc n)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   106
    using `increasing s` increasing_strict by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   107
  finally show ?thesis ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   108
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   109
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   110
lemma section_of2:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   111
  assumes "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   112
  shows "n < s (Suc (section_of s n))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   113
  unfolding section_of_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   114
  by (rule LeastI_ex) (rule section_help [OF `increasing s`])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   115
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   116
lemma section_of1:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   117
  assumes [simp, intro]: "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   118
  assumes "s i \<le> n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   119
  shows "s (section_of s n) \<le> n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   120
proof (rule classical)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   121
  let ?m = "section_of s n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   122
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   123
  assume "\<not> ?thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   124
  hence a: "n < s ?m" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   125
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   126
  have nonzero: "?m \<noteq> 0"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   127
  proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   128
    assume "?m = 0"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   129
    from increasing_weak have "s 0 \<le> s i" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   130
    also note `\<dots> \<le> n`
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   131
    finally show False using `?m = 0` `n < s ?m` by simp 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   132
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   133
  with a have "n < s (Suc (?m - 1))" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   134
  with Least_le have "?m \<le> ?m - 1"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   135
    unfolding section_of_def .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   136
  with nonzero show ?thesis by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   137
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   138
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   139
lemma section_of_known: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   140
  assumes [simp]: "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   141
  assumes in_sect: "k \<in> section s i"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   142
  shows "section_of s k = i" (is "?s = i")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   143
proof (rule classical)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   144
  assume "\<not> ?thesis"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   145
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   146
  hence "?s < i \<or> ?s > i" by arith
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   147
  thus ?thesis
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   148
  proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   149
    assume "?s < i"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   150
    hence "Suc ?s \<le> i" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   151
    with increasing_weak have "s (Suc ?s) \<le> s i" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   152
    moreover have "k < s (Suc ?s)" using section_of2 by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   153
    moreover from in_sect have "s i \<le> k" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   154
    ultimately show ?thesis by simp 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   155
  next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   156
    assume "i < ?s" hence "Suc i \<le> ?s" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   157
    with increasing_weak have "s (Suc i) \<le> s ?s" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   158
    moreover 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   159
    from in_sect have "s i \<le> k" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   160
    with section_of1 have "s ?s \<le> k" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   161
    moreover from in_sect have "k < s (Suc i)" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   162
    ultimately show ?thesis by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   163
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   164
qed 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   165
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   166
lemma in_section_of: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   167
  assumes "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   168
  assumes "s i \<le> k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   169
  shows "k \<in> section s (section_of s k)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   170
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   171
  by (auto intro:section_of1 section_of2)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   172
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   173
end