src/HOL/Relation_Power.thy
author haftmann
Tue, 19 May 2009 16:54:55 +0200
changeset 31205 98370b26c2ce
parent 30949 37f887b55e7f
permissions -rw-r--r--
String.literal replaces message_string, code_numeral replaces (code_)index
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Relation_Power.thy
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     2
    Author:     Tobias Nipkow
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     3
    Copyright   1996  TU Muenchen
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
     4
*)
11306
6f4ed75b2dca added comments
nipkow
parents: 11305
diff changeset
     5
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
     6
header{*Powers of Relations and Functions*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15112
diff changeset
     8
theory Relation_Power
29654
24e73987bfe2 Plain, Main form meeting points in import hierarchy
haftmann
parents: 28517
diff changeset
     9
imports Power Transitive_Closure Plain
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15112
diff changeset
    10
begin
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    11
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    12
consts funpower :: "('a \<Rightarrow> 'b) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'b" (infixr "^^" 80)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    13
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    14
overloading
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    15
  relpow \<equiv> "funpower \<Colon> ('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a \<times> 'a) set"
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    16
begin
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    17
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    18
text {* @{text "R ^^ n = R O ... O R"}, the n-fold composition of @{text R} *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    19
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    20
primrec relpow where
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    21
    "(R \<Colon> ('a \<times> 'a) set) ^^ 0 = Id"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    22
  | "(R \<Colon> ('a \<times> 'a) set) ^^ Suc n = R O (R ^^ n)"
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    23
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    24
end
11305
2ce86fccc95b added ^ on functions.
nipkow
parents: 10213
diff changeset
    25
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    26
overloading
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    27
  funpow \<equiv> "funpower \<Colon> ('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    28
begin
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    29
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    30
text {* @{text "f ^^ n = f o ... o f"}, the n-fold composition of @{text f} *}
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    31
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    32
primrec funpow where
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    33
    "(f \<Colon> 'a \<Rightarrow> 'a) ^^ 0 = id"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    34
  | "(f \<Colon> 'a \<Rightarrow> 'a) ^^ Suc n = f o (f ^^ n)"
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    35
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    36
end
11305
2ce86fccc95b added ^ on functions.
nipkow
parents: 10213
diff changeset
    37
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    38
primrec fun_pow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    39
    "fun_pow 0 f = id"
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    40
  | "fun_pow (Suc n) f = f o fun_pow n f"
21414
4cb808163adc workaround for definition violating type discipline
haftmann
parents: 20503
diff changeset
    41
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    42
lemma funpow_fun_pow [code unfold]:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    43
  "f ^^ n = fun_pow n f"
25861
494d9301cc75 refined overloading target
haftmann
parents: 25295
diff changeset
    44
  unfolding funpow_def fun_pow_def ..
21414
4cb808163adc workaround for definition violating type discipline
haftmann
parents: 20503
diff changeset
    45
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    46
lemma funpow_add:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    47
  "f ^^ (m + n) = f ^^ m o f ^^ n"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    48
  by (induct m) simp_all
15112
6f0772a94299 added a thm
nipkow
parents: 12338
diff changeset
    49
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    50
lemma funpow_swap1:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    51
  "f ((f ^^ n) x) = (f ^^ n) (f x)"
18049
156bba334c12 A few new lemmas
nipkow
parents: 15410
diff changeset
    52
proof -
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    53
  have "f ((f ^^ n) x) = (f ^^ (n+1)) x" unfolding One_nat_def by simp
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    54
  also have "\<dots>  = (f ^^ n o f ^^ 1) x" by (simp only: funpow_add)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    55
  also have "\<dots> = (f ^^ n) (f x)" unfolding One_nat_def by simp
18049
156bba334c12 A few new lemmas
nipkow
parents: 15410
diff changeset
    56
  finally show ?thesis .
156bba334c12 A few new lemmas
nipkow
parents: 15410
diff changeset
    57
qed
156bba334c12 A few new lemmas
nipkow
parents: 15410
diff changeset
    58
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    59
lemma rel_pow_1 [simp]:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    60
  fixes R :: "('a * 'a) set"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    61
  shows "R ^^ 1 = R"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    62
  by simp
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    63
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    64
lemma rel_pow_0_I: 
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    65
  "(x, x) \<in> R ^^ 0"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    66
  by simp
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    67
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    68
lemma rel_pow_Suc_I:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    69
  "(x, y) \<in>  R ^^ n \<Longrightarrow> (y, z) \<in> R \<Longrightarrow> (x, z) \<in> R ^^ Suc n"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    70
  by auto
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    71
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    72
lemma rel_pow_Suc_I2:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    73
  "(x, y) \<in> R \<Longrightarrow> (y, z) \<in> R ^^ n \<Longrightarrow> (x, z) \<in> R ^^ Suc n"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    74
  by (induct n arbitrary: z) (simp, fastsimp)
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    75
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    76
lemma rel_pow_0_E:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    77
  "(x, y) \<in> R ^^ 0 \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    78
  by simp
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    79
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    80
lemma rel_pow_Suc_E:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    81
  "(x, z) \<in> R ^^ Suc n \<Longrightarrow> (\<And>y. (x, y) \<in> R ^^ n \<Longrightarrow> (y, z) \<in> R \<Longrightarrow> P) \<Longrightarrow> P"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    82
  by auto
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    83
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    84
lemma rel_pow_E:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    85
  "(x, z) \<in>  R ^^ n \<Longrightarrow>  (n = 0 \<Longrightarrow> x = z \<Longrightarrow> P)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    86
   \<Longrightarrow> (\<And>y m. n = Suc m \<Longrightarrow> (x, y) \<in>  R ^^ m \<Longrightarrow> (y, z) \<in> R \<Longrightarrow> P)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    87
   \<Longrightarrow> P"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    88
  by (cases n) auto
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    89
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    90
lemma rel_pow_Suc_D2:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    91
  "(x, z) \<in> R ^^ Suc n \<Longrightarrow> (\<exists>y. (x, y) \<in> R \<and> (y, z) \<in> R ^^ n)"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 18398
diff changeset
    92
  apply (induct n arbitrary: x z)
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    93
   apply (blast intro: rel_pow_0_I elim: rel_pow_0_E rel_pow_Suc_E)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    94
  apply (blast intro: rel_pow_Suc_I elim: rel_pow_0_E rel_pow_Suc_E)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    95
  done
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    96
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
    97
lemma rel_pow_Suc_D2':
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
    98
  "\<forall>x y z. (x, y) \<in> R ^^ n \<and> (y, z) \<in> R \<longrightarrow> (\<exists>w. (x, w) \<in> R \<and> (w, z) \<in> R ^^ n)"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
    99
  by (induct n) (simp_all, blast)
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   100
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   101
lemma rel_pow_E2:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   102
  "(x, z) \<in> R ^^ n \<Longrightarrow>  (n = 0 \<Longrightarrow> x = z \<Longrightarrow> P)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   103
     \<Longrightarrow> (\<And>y m. n = Suc m \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (y, z) \<in> R ^^ m \<Longrightarrow> P)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   104
   \<Longrightarrow> P"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   105
  apply (cases n, simp)
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   106
  apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   107
  done
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   108
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   109
lemma rtrancl_imp_UN_rel_pow:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   110
  "p \<in> R^* \<Longrightarrow> p \<in> (\<Union>n. R ^^ n)"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   111
  apply (cases p) apply (simp only:)
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   112
  apply (erule rtrancl_induct)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   113
   apply (blast intro: rel_pow_0_I rel_pow_Suc_I)+
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   114
  done
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   115
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   116
lemma rel_pow_imp_rtrancl:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   117
  "p \<in> R ^^ n \<Longrightarrow> p \<in> R^*"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   118
  apply (induct n arbitrary: p)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   119
  apply (simp_all only: split_tupled_all)
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   120
   apply (blast intro: rtrancl_refl elim: rel_pow_0_E)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   121
  apply (blast elim: rel_pow_Suc_E intro: rtrancl_into_rtrancl)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   122
  done
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   123
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   124
lemma rtrancl_is_UN_rel_pow:
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   125
  "R^* = (UN n. R ^^ n)"
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   126
  by (blast intro: rtrancl_imp_UN_rel_pow rel_pow_imp_rtrancl)
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   127
25295
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   128
lemma trancl_power:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   129
  "x \<in> r^+ = (\<exists>n > 0. x \<in> r ^^ n)"
25295
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   130
  apply (cases x)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   131
  apply simp
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   132
  apply (rule iffI)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   133
   apply (drule tranclD2)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   134
   apply (clarsimp simp: rtrancl_is_UN_rel_pow)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   135
   apply (rule_tac x="Suc x" in exI)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   136
   apply (clarsimp simp: rel_comp_def)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   137
   apply fastsimp
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   138
  apply clarsimp
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   139
  apply (case_tac n, simp)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   140
  apply clarsimp
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   141
  apply (drule rel_pow_imp_rtrancl)
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   142
  apply fastsimp
12985023be5e tranclD2 (tranclD at the other end) + trancl_power
kleing
parents: 24996
diff changeset
   143
  done
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   144
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   145
lemma single_valued_rel_pow:
30949
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   146
  fixes R :: "('a * 'a) set"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   147
  shows "single_valued R \<Longrightarrow> single_valued (R ^^ n)"
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   148
  apply (induct n arbitrary: R)
37f887b55e7f separated funpow, relpow from power on monoids
haftmann
parents: 30079
diff changeset
   149
  apply simp_all
18398
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   150
  apply (rule single_valuedI)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   151
  apply (fast dest: single_valuedD elim: rel_pow_Suc_E)
5d63a8b35688 tuned proofs;
wenzelm
parents: 18049
diff changeset
   152
  done
15410
18914688a5fd converted Relation_Power to new-style theory
paulson
parents: 15140
diff changeset
   153
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   154
end