10213
|
1 |
(* Title: HOL/Relation_Power.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow
|
|
4 |
Copyright 1996 TU Muenchen
|
15410
|
5 |
*)
|
11306
|
6 |
|
15410
|
7 |
header{*Powers of Relations and Functions*}
|
10213
|
8 |
|
15131
|
9 |
theory Relation_Power
|
15140
|
10 |
imports Nat
|
15131
|
11 |
begin
|
10213
|
12 |
|
|
13 |
instance
|
18398
|
14 |
set :: (type) power ..
|
15410
|
15 |
--{* only type @{typ "('a * 'a) set"} should be in class @{text power}!*}
|
10213
|
16 |
|
15410
|
17 |
(*R^n = R O ... O R, the n-fold composition of R*)
|
10213
|
18 |
primrec (relpow)
|
|
19 |
"R^0 = Id"
|
|
20 |
"R^(Suc n) = R O (R^n)"
|
|
21 |
|
11305
|
22 |
|
11306
|
23 |
instance
|
15410
|
24 |
fun :: (type, type) power ..
|
|
25 |
--{* only type @{typ "'a => 'a"} should be in class @{text power}!*}
|
11305
|
26 |
|
15410
|
27 |
(*f^n = f o ... o f, the n-fold composition of f*)
|
11305
|
28 |
primrec (funpow)
|
|
29 |
"f^0 = id"
|
|
30 |
"f^(Suc n) = f o (f^n)"
|
|
31 |
|
15410
|
32 |
text{*WARNING: due to the limits of Isabelle's type classes, exponentiation on
|
|
33 |
functions and relations has too general a domain, namely @{typ "('a * 'b)set"}
|
|
34 |
and @{typ "'a => 'b"}. Explicit type constraints may therefore be necessary.
|
|
35 |
For example, @{term "range(f^n) = A"} and @{term "Range(R^n) = B"} need
|
|
36 |
constraints.*}
|
|
37 |
|
15112
|
38 |
lemma funpow_add: "f ^ (m+n) = f^m o f^n"
|
18398
|
39 |
by (induct m) simp_all
|
15112
|
40 |
|
18049
|
41 |
lemma funpow_swap1: "f((f^n) x) = (f^n)(f x)"
|
|
42 |
proof -
|
|
43 |
have "f((f^n) x) = (f^(n+1)) x" by simp
|
18398
|
44 |
also have "\<dots> = (f^n o f^1) x" by (simp only: funpow_add)
|
18049
|
45 |
also have "\<dots> = (f^n)(f x)" by simp
|
|
46 |
finally show ?thesis .
|
|
47 |
qed
|
|
48 |
|
18398
|
49 |
lemma rel_pow_1 [simp]:
|
|
50 |
fixes R :: "('a*'a)set"
|
|
51 |
shows "R^1 = R"
|
|
52 |
by simp
|
15410
|
53 |
|
|
54 |
lemma rel_pow_0_I: "(x,x) : R^0"
|
18398
|
55 |
by simp
|
15410
|
56 |
|
|
57 |
lemma rel_pow_Suc_I: "[| (x,y) : R^n; (y,z):R |] ==> (x,z):R^(Suc n)"
|
18398
|
58 |
by auto
|
15410
|
59 |
|
18398
|
60 |
lemma rel_pow_Suc_I2:
|
|
61 |
"(x, y) : R \<Longrightarrow> (y, z) : R^n \<Longrightarrow> (x,z) : R^(Suc n)"
|
|
62 |
apply (induct n fixing: z)
|
|
63 |
apply simp
|
|
64 |
apply fastsimp
|
|
65 |
done
|
15410
|
66 |
|
|
67 |
lemma rel_pow_0_E: "[| (x,y) : R^0; x=y ==> P |] ==> P"
|
18398
|
68 |
by simp
|
15410
|
69 |
|
18398
|
70 |
lemma rel_pow_Suc_E:
|
|
71 |
"[| (x,z) : R^(Suc n); !!y. [| (x,y) : R^n; (y,z) : R |] ==> P |] ==> P"
|
|
72 |
by auto
|
15410
|
73 |
|
18398
|
74 |
lemma rel_pow_E:
|
|
75 |
"[| (x,z) : R^n; [| n=0; x = z |] ==> P;
|
|
76 |
!!y m. [| n = Suc m; (x,y) : R^m; (y,z) : R |] ==> P
|
15410
|
77 |
|] ==> P"
|
18398
|
78 |
by (cases n) auto
|
15410
|
79 |
|
18398
|
80 |
lemma rel_pow_Suc_D2:
|
|
81 |
"(x, z) : R^(Suc n) \<Longrightarrow> (\<exists>y. (x,y) : R & (y,z) : R^n)"
|
|
82 |
apply (induct n fixing: x z)
|
|
83 |
apply (blast intro: rel_pow_0_I elim: rel_pow_0_E rel_pow_Suc_E)
|
|
84 |
apply (blast intro: rel_pow_Suc_I elim: rel_pow_0_E rel_pow_Suc_E)
|
|
85 |
done
|
15410
|
86 |
|
|
87 |
lemma rel_pow_Suc_D2':
|
18398
|
88 |
"\<forall>x y z. (x,y) : R^n & (y,z) : R --> (\<exists>w. (x,w) : R & (w,z) : R^n)"
|
|
89 |
by (induct n) (simp_all, blast)
|
15410
|
90 |
|
18398
|
91 |
lemma rel_pow_E2:
|
|
92 |
"[| (x,z) : R^n; [| n=0; x = z |] ==> P;
|
|
93 |
!!y m. [| n = Suc m; (x,y) : R; (y,z) : R^m |] ==> P
|
15410
|
94 |
|] ==> P"
|
18398
|
95 |
apply (case_tac n, simp)
|
|
96 |
apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast)
|
|
97 |
done
|
15410
|
98 |
|
|
99 |
lemma rtrancl_imp_UN_rel_pow: "!!p. p:R^* ==> p : (UN n. R^n)"
|
18398
|
100 |
apply (simp only: split_tupled_all)
|
|
101 |
apply (erule rtrancl_induct)
|
|
102 |
apply (blast intro: rel_pow_0_I rel_pow_Suc_I)+
|
|
103 |
done
|
15410
|
104 |
|
|
105 |
lemma rel_pow_imp_rtrancl: "!!p. p:R^n ==> p:R^*"
|
18398
|
106 |
apply (simp only: split_tupled_all)
|
|
107 |
apply (induct n)
|
|
108 |
apply (blast intro: rtrancl_refl elim: rel_pow_0_E)
|
|
109 |
apply (blast elim: rel_pow_Suc_E intro: rtrancl_into_rtrancl)
|
|
110 |
done
|
15410
|
111 |
|
|
112 |
lemma rtrancl_is_UN_rel_pow: "R^* = (UN n. R^n)"
|
18398
|
113 |
by (blast intro: rtrancl_imp_UN_rel_pow rel_pow_imp_rtrancl)
|
15410
|
114 |
|
|
115 |
|
18398
|
116 |
lemma single_valued_rel_pow:
|
|
117 |
"!!r::('a * 'a)set. single_valued r ==> single_valued (r^n)"
|
|
118 |
apply (rule single_valuedI)
|
|
119 |
apply (induct n)
|
|
120 |
apply simp
|
|
121 |
apply (fast dest: single_valuedD elim: rel_pow_Suc_E)
|
|
122 |
done
|
15410
|
123 |
|
|
124 |
ML
|
|
125 |
{*
|
|
126 |
val funpow_add = thm "funpow_add";
|
|
127 |
val rel_pow_1 = thm "rel_pow_1";
|
|
128 |
val rel_pow_0_I = thm "rel_pow_0_I";
|
|
129 |
val rel_pow_Suc_I = thm "rel_pow_Suc_I";
|
|
130 |
val rel_pow_Suc_I2 = thm "rel_pow_Suc_I2";
|
|
131 |
val rel_pow_0_E = thm "rel_pow_0_E";
|
|
132 |
val rel_pow_Suc_E = thm "rel_pow_Suc_E";
|
|
133 |
val rel_pow_E = thm "rel_pow_E";
|
|
134 |
val rel_pow_Suc_D2 = thm "rel_pow_Suc_D2";
|
|
135 |
val rel_pow_Suc_D2 = thm "rel_pow_Suc_D2";
|
|
136 |
val rel_pow_E2 = thm "rel_pow_E2";
|
|
137 |
val rtrancl_imp_UN_rel_pow = thm "rtrancl_imp_UN_rel_pow";
|
|
138 |
val rel_pow_imp_rtrancl = thm "rel_pow_imp_rtrancl";
|
|
139 |
val rtrancl_is_UN_rel_pow = thm "rtrancl_is_UN_rel_pow";
|
|
140 |
val single_valued_rel_pow = thm "single_valued_rel_pow";
|
|
141 |
*}
|
|
142 |
|
10213
|
143 |
end
|