src/HOL/SizeChange/Kleene_Algebras.thy
author haftmann
Tue, 22 Apr 2008 22:00:25 +0200
changeset 26739 947b6013e863
parent 25314 5eaf3e8b50a4
child 27682 25aceefd4786
permissions -rw-r--r--
different handling of eq class for nbe
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/Kleene_Algebras.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 "Kleene Algebras"
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 Kleene_Algebras
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
text {* A type class of kleene algebras *}
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
class star = type +
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    15
  fixes star :: "'a \<Rightarrow> 'a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    16
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    17
class idem_add = ab_semigroup_add +
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    18
  assumes add_idem [simp]: "x + x = x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    19
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    20
lemma add_idem2[simp]: "(x::'a::idem_add) + (x + y) = x + y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    21
  unfolding add_assoc[symmetric]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    22
  by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    23
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    24
class order_by_add = idem_add + ord +
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    25
  assumes order_def: "a \<le> b \<longleftrightarrow> a + b = b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    26
  assumes strict_order_def: "a < b \<longleftrightarrow> a \<le> b \<and> a \<noteq> b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    27
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    28
lemma ord_simp1[simp]: "(x::'a::order_by_add) \<le> y \<Longrightarrow> x + y = y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    29
  unfolding order_def .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    30
lemma ord_simp2[simp]: "(x::'a::order_by_add) \<le> y \<Longrightarrow> y + x = y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    31
  unfolding order_def add_commute .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    32
lemma ord_intro: "(x::'a::order_by_add) + y = y \<Longrightarrow> x \<le> y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    33
  unfolding order_def .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    34
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    35
instance order_by_add \<subseteq> order
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    36
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    37
  fix x y z :: 'a
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    38
  show "x \<le> x" unfolding order_def by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    39
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    40
  show "\<lbrakk>x \<le> y; y \<le> z\<rbrakk> \<Longrightarrow> x \<le> z"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    41
  proof (rule ord_intro)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    42
    assume "x \<le> y" "y \<le> z"
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
    have "x + z = x + y + z" by (simp add:`y \<le> z` add_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    45
    also have "\<dots> = y + z" by (simp add:`x \<le> y`)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    46
    also have "\<dots> = z" by (simp add:`y \<le> z`)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    47
    finally show "x + z = z" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    48
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    49
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    50
  show "\<lbrakk>x \<le> y; y \<le> x\<rbrakk> \<Longrightarrow> x = y" unfolding order_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    51
    by (simp add:add_commute)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    52
  show "x < y \<longleftrightarrow> x \<le> y \<and> x \<noteq> y" by (fact strict_order_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    53
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    54
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    55
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    56
class pre_kleene = semiring_1 + order_by_add
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
instance pre_kleene \<subseteq> pordered_semiring
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    59
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    60
  fix x y z :: 'a
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    61
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    62
  assume "x \<le> y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    63
   
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    64
  show "z + x \<le> z + y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    65
  proof (rule ord_intro)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    66
    have "z + x + (z + y) = x + y + z" by (simp add:add_ac)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    67
    also have "\<dots> = z + y" by (simp add:`x \<le> y` add_ac)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    68
    finally show "z + x + (z + y) = z + y" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    69
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    70
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    71
  show "z * x \<le> z * y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    72
  proof (rule ord_intro)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    73
    from `x \<le> y` have "z * (x + y) = z * y" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    74
    thus "z * x + z * y = z * y" by (simp add:right_distrib)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    75
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    76
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    77
  show "x * z \<le> y * z"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    78
  proof (rule ord_intro)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    79
    from `x \<le> y` have "(x + y) * z = y * z" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    80
    thus "x * z + y * z = y * z" by (simp add:left_distrib)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    81
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    82
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    83
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    84
class kleene = pre_kleene + star +
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    85
  assumes star1: "1 + a * star a \<le> star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    86
  and star2: "1 + star a * a \<le> star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    87
  and star3: "a * x \<le> x \<Longrightarrow> star a * x \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    88
  and star4: "x * a \<le> x \<Longrightarrow> x * star a \<le> x"
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
class kleene_by_complete_lattice = pre_kleene
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    91
  + complete_lattice + recpower + star +
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    92
  assumes star_cont: "a * star b * c = SUPR UNIV (\<lambda>n. a * b ^ n * c)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    93
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    94
lemma plus_leI: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    95
  fixes x :: "'a :: order_by_add"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    96
  shows "x \<le> z \<Longrightarrow> y \<le> z \<Longrightarrow> x + y \<le> z"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    97
  unfolding order_def by (simp add:add_assoc)
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 le_SUPI':
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   100
  fixes l :: "'a :: complete_lattice"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   101
  assumes "l \<le> M i"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   102
  shows "l \<le> (SUP i. M i)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   103
  using assms by (rule order_trans) (rule le_SUPI [OF UNIV_I])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   104
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   105
lemma zero_minimum[simp]: "(0::'a::pre_kleene) \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   106
  unfolding order_def by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   107
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   108
instance kleene_by_complete_lattice \<subseteq> kleene
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   109
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   110
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   111
  fix a x :: 'a
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   112
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   113
  have [simp]: "1 \<le> star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   114
    unfolding star_cont[of 1 a 1, simplified] 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   115
    by (subst power_0[symmetric]) (rule le_SUPI [OF UNIV_I])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   116
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   117
  show "1 + a * star a \<le> star a" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   118
    apply (rule plus_leI, simp)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   119
    apply (simp add:star_cont[of a a 1, simplified])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   120
    apply (simp add:star_cont[of 1 a 1, simplified])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   121
    apply (subst power_Suc[symmetric])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   122
    by (intro SUP_leI le_SUPI UNIV_I)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   123
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   124
  show "1 + star a * a \<le> star a" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   125
    apply (rule plus_leI, simp)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   126
    apply (simp add:star_cont[of 1 a a, simplified])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   127
    apply (simp add:star_cont[of 1 a 1, simplified])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   128
    by (auto intro: SUP_leI le_SUPI UNIV_I simp add: power_Suc[symmetric] power_commutes)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   129
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   130
  show "a * x \<le> x \<Longrightarrow> star a * x \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   131
  proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   132
    assume a: "a * x \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   133
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   134
    {
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   135
      fix n
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   136
      have "a ^ (Suc n) * x \<le> a ^ n * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   137
      proof (induct n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   138
        case 0 thus ?case by (simp add:a power_Suc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   139
      next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   140
        case (Suc n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   141
        hence "a * (a ^ Suc n * x) \<le> a * (a ^ n * x)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   142
          by (auto intro: mult_mono)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   143
        thus ?case
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   144
          by (simp add:power_Suc mult_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   145
      qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   146
    }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   147
    note a = this
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   148
    
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   149
    {
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   150
      fix n have "a ^ n * x \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   151
      proof (induct n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   152
        case 0 show ?case by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   153
      next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   154
        case (Suc n) with a[of n]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   155
        show ?case by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   156
      qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   157
    }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   158
    note b = this
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   159
    
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   160
    show "star a * x \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   161
      unfolding star_cont[of 1 a x, simplified]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   162
      by (rule SUP_leI) (rule b)
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
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   165
  show "x * a \<le> x \<Longrightarrow> x * star a \<le> x" (* symmetric *)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   166
  proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   167
    assume a: "x * a \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   168
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   169
    {
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   170
      fix n
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   171
      have "x * a ^ (Suc n) \<le> x * a ^ n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   172
      proof (induct n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   173
        case 0 thus ?case by (simp add:a power_Suc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   174
      next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   175
        case (Suc n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   176
        hence "(x * a ^ Suc n) * a  \<le> (x * a ^ n) * a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   177
          by (auto intro: mult_mono)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   178
        thus ?case
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   179
          by (simp add:power_Suc power_commutes mult_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   180
      qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   181
    }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   182
    note a = this
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   183
    
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   184
    {
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   185
      fix n have "x * a ^ n \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   186
      proof (induct n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   187
        case 0 show ?case by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   188
      next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   189
        case (Suc n) with a[of n]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   190
        show ?case by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   191
      qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   192
    }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   193
    note b = this
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   194
    
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   195
    show "x * star a \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   196
      unfolding star_cont[of x a 1, simplified]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   197
      by (rule SUP_leI) (rule b)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   198
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   199
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   200
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   201
lemma less_add[simp]:  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   202
  fixes a b :: "'a :: order_by_add"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   203
  shows "a \<le> a + b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   204
  and "b \<le> a + b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   205
  unfolding order_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   206
  by (auto simp:add_ac)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   207
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   208
lemma add_est1:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   209
  fixes a b c :: "'a :: order_by_add"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   210
  assumes a: "a + b \<le> c"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   211
  shows "a \<le> c"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   212
  using less_add(1) a
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   213
  by (rule order_trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   214
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   215
lemma add_est2:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   216
  fixes a b c :: "'a :: order_by_add"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   217
  assumes a: "a + b \<le> c"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   218
  shows "b \<le> c"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   219
  using less_add(2) a
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   220
  by (rule order_trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   221
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   222
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   223
lemma star3':
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   224
  fixes a b x :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   225
  assumes a: "b + a * x \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   226
  shows "star a * b \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   227
proof (rule order_trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   228
  from a have "b \<le> x" by (rule add_est1)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   229
  show "star a * b \<le> star a * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   230
    by (rule mult_mono) (auto simp:`b \<le> x`)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   231
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   232
  from a have "a * x \<le> x" by (rule add_est2)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   233
  with star3 show "star a * x \<le> x" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   234
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   235
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   236
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   237
lemma star4':
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   238
  fixes a b x :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   239
  assumes a: "b + x * a \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   240
  shows "b * star a \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   241
proof (rule order_trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   242
  from a have "b \<le> x" by (rule add_est1)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   243
  show "b * star a \<le> x * star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   244
    by (rule mult_mono) (auto simp:`b \<le> x`)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   245
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   246
  from a have "x * a \<le> x" by (rule add_est2)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   247
  with star4 show "x * star a \<le> x" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   248
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   249
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   250
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   251
lemma star_idemp:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   252
  fixes x :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   253
  shows "star (star x) = star x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   254
  oops
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   255
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   256
lemma star_unfold_left:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   257
  fixes a :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   258
  shows "1 + a * star a = star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   259
proof (rule order_antisym, rule star1)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   260
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   261
  have "1 + a * (1 + a * star a) \<le> 1 + a * star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   262
    apply (rule add_mono, rule)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   263
    apply (rule mult_mono, auto)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   264
    apply (rule star1)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   265
    done
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   266
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   267
  with star3' have "star a * 1 \<le> 1 + a * star a" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   268
  thus "star a \<le> 1 + a * star a" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   269
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   270
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   271
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   272
lemma star_unfold_right:  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   273
  fixes a :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   274
  shows "1 + star a * a = star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   275
proof (rule order_antisym, rule star2)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   276
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   277
  have "1 + (1 + star a * a) * a \<le> 1 + star a * a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   278
    apply (rule add_mono, rule)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   279
    apply (rule mult_mono, auto)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   280
    apply (rule star2)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   281
    done
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   282
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   283
  with star4' have "1 * star a \<le> 1 + star a * a" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   284
  thus "star a \<le> 1 + star a * a" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   285
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   286
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   287
lemma star_zero[simp]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   288
  shows "star (0::'a::kleene) = 1"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   289
  by (rule star_unfold_left[of 0, simplified])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   290
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   291
lemma star_commute:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   292
  fixes a b x :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   293
  assumes a: "a * x = x * b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   294
  shows "star a * x = x * star b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   295
proof (rule order_antisym)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   296
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   297
  show "star a * x \<le> x * star b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   298
  proof (rule star3', rule order_trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   299
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   300
    from a have "a * x \<le> x * b" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   301
    hence "a * x * star b \<le> x * b * star b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   302
      by (rule mult_mono) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   303
    thus "x + a * (x * star b) \<le> x + x * b * star b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   304
      using add_mono by (auto simp: mult_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   305
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   306
    show "\<dots> \<le> x * star b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   307
    proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   308
      have "x * (1 + b * star b) \<le> x * star b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   309
        by (rule mult_mono[OF _ star1]) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   310
      thus ?thesis
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   311
        by (simp add:right_distrib mult_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   312
    qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   313
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   315
  show "x * star b \<le> star a * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   316
  proof (rule star4', rule order_trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   317
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   318
    from a have b: "x * b \<le> a * x" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   319
    have "star a * x * b \<le> star a * a * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   320
      unfolding mult_assoc
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   321
      by (rule mult_mono[OF _ b]) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   322
    thus "x + star a * x * b \<le> x + star a * a * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   323
      using add_mono by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   324
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   325
    show "\<dots> \<le> star a * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   326
    proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   327
      have "(1 + star a * a) * x \<le> star a * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   328
        by (rule mult_mono[OF star2]) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   329
      thus ?thesis
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   330
        by (simp add:left_distrib mult_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   331
    qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   332
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   333
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   334
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   335
lemma star_assoc:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   336
  fixes c d :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   337
  shows "star (c * d) * c = c * star (d * c)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   338
  by (auto simp:mult_assoc star_commute)  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   339
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   340
lemma star_dist:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   341
  fixes a b :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   342
  shows "star (a + b) = star a * star (b * star a)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   343
  oops
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   344
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   345
lemma star_one:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   346
  fixes a p p' :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   347
  assumes "p * p' = 1" and "p' * p = 1"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   348
  shows "p' * star a * p = star (p' * a * p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   349
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   350
  from assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   351
  have "p' * star a * p = p' * star (p * p' * a) * p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   352
    by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   353
  also have "\<dots> = p' * p * star (p' * a * p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   354
    by (simp add: mult_assoc star_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   355
  also have "\<dots> = star (p' * a * p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   356
    by (simp add: assms)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   357
  finally show ?thesis .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   358
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   359
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   360
lemma star_mono:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   361
  fixes x y :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   362
  assumes "x \<le> y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   363
  shows "star x \<le> star y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   364
  oops
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   365
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   366
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   367
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   368
(* Own lemmas *)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   369
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   370
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   371
lemma x_less_star[simp]: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   372
  fixes x :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   373
  shows "x \<le> x * star a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   374
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   375
  have "x \<le> x * (1 + a * star a)" by (simp add:right_distrib)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   376
  also have "\<dots> = x * star a" by (simp only: star_unfold_left)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   377
  finally show ?thesis .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   378
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   379
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   380
subsection {* Transitive Closure *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   381
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   382
definition 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   383
  "tcl (x::'a::kleene) = star x * x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   384
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   385
lemma tcl_zero: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   386
  "tcl (0::'a::kleene) = 0"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   387
  unfolding tcl_def by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   388
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   389
lemma tcl_unfold_right: "tcl a = a + tcl a * a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   390
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   391
  from star_unfold_right[of a]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   392
  have "a * (1 + star a * a) = a * star a" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   393
  from this[simplified right_distrib, simplified]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   394
  show ?thesis
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   395
    by (simp add:tcl_def star_commute mult_ac)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   396
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   397
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   398
lemma less_tcl: "a \<le> tcl a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   399
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   400
  have "a \<le> a + tcl a * a" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   401
  also have "\<dots> = tcl a" by (rule tcl_unfold_right[symmetric])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   402
  finally show ?thesis .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   403
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   404
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   405
subsection {* Naive Algorithm to generate the transitive closure *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   406
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   407
function (default "\<lambda>x. 0", tailrec, domintros)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   408
  mk_tcl :: "('a::{plus,times,ord,zero}) \<Rightarrow> 'a \<Rightarrow> 'a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   409
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   410
  "mk_tcl A X = (if X * A \<le> X then X else mk_tcl A (X + X * A))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   411
  by pat_completeness simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   412
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   413
declare mk_tcl.simps[simp del] (* loops *)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   414
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   415
lemma mk_tcl_code[code]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   416
  "mk_tcl A X = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   417
  (let XA = X * A 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   418
  in if XA \<le> X then X else mk_tcl A (X + XA))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   419
  unfolding mk_tcl.simps[of A X] Let_def ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   420
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   421
lemma mk_tcl_lemma1:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   422
  fixes X :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   423
  shows "(X + X * A) * star A = X * star A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   424
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   425
  have "A * star A \<le> 1 + A * star A" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   426
  also have "\<dots> = star A" by (simp add:star_unfold_left)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   427
  finally have "star A + A * star A = star A" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   428
  hence "X * (star A + A * star A) = X * star A" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   429
  thus ?thesis by (simp add:left_distrib right_distrib mult_ac)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   430
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   431
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   432
lemma mk_tcl_lemma2:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   433
  fixes X :: "'a :: kleene"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   434
  shows "X * A \<le> X \<Longrightarrow> X * star A = X"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   435
  by (rule order_antisym) (auto simp:star4)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   436
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   437
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   438
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   439
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   440
lemma mk_tcl_correctness:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   441
  fixes A X :: "'a :: {kleene}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   442
  assumes "mk_tcl_dom (A, X)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   443
  shows "mk_tcl A X = X * star A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   444
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   445
  by induct (auto simp:mk_tcl_lemma1 mk_tcl_lemma2)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   446
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   447
lemma graph_implies_dom: "mk_tcl_graph x y \<Longrightarrow> mk_tcl_dom x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   448
  by (rule mk_tcl_graph.induct) (auto intro:accp.accI elim:mk_tcl_rel.cases)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   449
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   450
lemma mk_tcl_default: "\<not> mk_tcl_dom (a,x) \<Longrightarrow> mk_tcl a x = 0"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   451
  unfolding mk_tcl_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   452
  by (rule fundef_default_value[OF mk_tcl_sum_def graph_implies_dom])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   453
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   454
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   455
text {* We can replace the dom-Condition of the correctness theorem 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   456
  with something executable *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   457
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   458
lemma mk_tcl_correctness2:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   459
  fixes A X :: "'a :: {kleene}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   460
  assumes "mk_tcl A A \<noteq> 0"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   461
  shows "mk_tcl A A = tcl A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   462
  using assms mk_tcl_default mk_tcl_correctness
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   463
  unfolding tcl_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   464
  by (auto simp:star_commute)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   465
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   466
end