author | wenzelm |
Sat, 18 Aug 2007 13:32:28 +0200 | |
changeset 24325 | 5c29e8822f50 |
parent 24304 | 69d40a562ba4 |
child 24996 | ebd5f4cc7118 |
permissions | -rw-r--r-- |
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*) |
24304 | 18 |
primrec (unchecked relpow) |
10213 | 19 |
"R^0 = Id" |
20 |
"R^(Suc n) = R O (R^n)" |
|
21 |
||
11305 | 22 |
|
11306 | 23 |
instance |
21414
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
24 |
"fun" :: (type, type) power .. |
15410 | 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*) |
24304 | 28 |
primrec (unchecked funpow) |
11305 | 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 |
||
21414
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
38 |
text {* |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
39 |
Circumvent this problem for code generation: |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
40 |
*} |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
41 |
|
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
42 |
definition |
22737 | 43 |
funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" |
44 |
where |
|
21414
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
45 |
funpow_def: "funpow n f = f ^ n" |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
46 |
|
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
47 |
lemmas [code inline] = funpow_def [symmetric] |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
48 |
|
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
49 |
lemma [code func]: |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
50 |
"funpow 0 f = id" |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
51 |
"funpow (Suc n) f = f o funpow n f" |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
52 |
unfolding funpow_def by simp_all |
4cb808163adc
workaround for definition violating type discipline
haftmann
parents:
20503
diff
changeset
|
53 |
|
15112 | 54 |
lemma funpow_add: "f ^ (m+n) = f^m o f^n" |
18398 | 55 |
by (induct m) simp_all |
15112 | 56 |
|
18049 | 57 |
lemma funpow_swap1: "f((f^n) x) = (f^n)(f x)" |
58 |
proof - |
|
59 |
have "f((f^n) x) = (f^(n+1)) x" by simp |
|
18398 | 60 |
also have "\<dots> = (f^n o f^1) x" by (simp only: funpow_add) |
18049 | 61 |
also have "\<dots> = (f^n)(f x)" by simp |
62 |
finally show ?thesis . |
|
63 |
qed |
|
64 |
||
18398 | 65 |
lemma rel_pow_1 [simp]: |
66 |
fixes R :: "('a*'a)set" |
|
67 |
shows "R^1 = R" |
|
68 |
by simp |
|
15410 | 69 |
|
70 |
lemma rel_pow_0_I: "(x,x) : R^0" |
|
18398 | 71 |
by simp |
15410 | 72 |
|
73 |
lemma rel_pow_Suc_I: "[| (x,y) : R^n; (y,z):R |] ==> (x,z):R^(Suc n)" |
|
18398 | 74 |
by auto |
15410 | 75 |
|
18398 | 76 |
lemma rel_pow_Suc_I2: |
77 |
"(x, y) : R \<Longrightarrow> (y, z) : R^n \<Longrightarrow> (x,z) : R^(Suc n)" |
|
20503 | 78 |
apply (induct n arbitrary: z) |
18398 | 79 |
apply simp |
80 |
apply fastsimp |
|
81 |
done |
|
15410 | 82 |
|
83 |
lemma rel_pow_0_E: "[| (x,y) : R^0; x=y ==> P |] ==> P" |
|
18398 | 84 |
by simp |
15410 | 85 |
|
18398 | 86 |
lemma rel_pow_Suc_E: |
87 |
"[| (x,z) : R^(Suc n); !!y. [| (x,y) : R^n; (y,z) : R |] ==> P |] ==> P" |
|
88 |
by auto |
|
15410 | 89 |
|
18398 | 90 |
lemma rel_pow_E: |
91 |
"[| (x,z) : R^n; [| n=0; x = z |] ==> P; |
|
92 |
!!y m. [| n = Suc m; (x,y) : R^m; (y,z) : R |] ==> P |
|
15410 | 93 |
|] ==> P" |
18398 | 94 |
by (cases n) auto |
15410 | 95 |
|
18398 | 96 |
lemma rel_pow_Suc_D2: |
97 |
"(x, z) : R^(Suc n) \<Longrightarrow> (\<exists>y. (x,y) : R & (y,z) : R^n)" |
|
20503 | 98 |
apply (induct n arbitrary: x z) |
18398 | 99 |
apply (blast intro: rel_pow_0_I elim: rel_pow_0_E rel_pow_Suc_E) |
100 |
apply (blast intro: rel_pow_Suc_I elim: rel_pow_0_E rel_pow_Suc_E) |
|
101 |
done |
|
15410 | 102 |
|
103 |
lemma rel_pow_Suc_D2': |
|
18398 | 104 |
"\<forall>x y z. (x,y) : R^n & (y,z) : R --> (\<exists>w. (x,w) : R & (w,z) : R^n)" |
105 |
by (induct n) (simp_all, blast) |
|
15410 | 106 |
|
18398 | 107 |
lemma rel_pow_E2: |
108 |
"[| (x,z) : R^n; [| n=0; x = z |] ==> P; |
|
109 |
!!y m. [| n = Suc m; (x,y) : R; (y,z) : R^m |] ==> P |
|
15410 | 110 |
|] ==> P" |
18398 | 111 |
apply (case_tac n, simp) |
112 |
apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast) |
|
113 |
done |
|
15410 | 114 |
|
115 |
lemma rtrancl_imp_UN_rel_pow: "!!p. p:R^* ==> p : (UN n. R^n)" |
|
18398 | 116 |
apply (simp only: split_tupled_all) |
117 |
apply (erule rtrancl_induct) |
|
118 |
apply (blast intro: rel_pow_0_I rel_pow_Suc_I)+ |
|
119 |
done |
|
15410 | 120 |
|
121 |
lemma rel_pow_imp_rtrancl: "!!p. p:R^n ==> p:R^*" |
|
18398 | 122 |
apply (simp only: split_tupled_all) |
123 |
apply (induct n) |
|
124 |
apply (blast intro: rtrancl_refl elim: rel_pow_0_E) |
|
125 |
apply (blast elim: rel_pow_Suc_E intro: rtrancl_into_rtrancl) |
|
126 |
done |
|
15410 | 127 |
|
128 |
lemma rtrancl_is_UN_rel_pow: "R^* = (UN n. R^n)" |
|
18398 | 129 |
by (blast intro: rtrancl_imp_UN_rel_pow rel_pow_imp_rtrancl) |
15410 | 130 |
|
131 |
||
18398 | 132 |
lemma single_valued_rel_pow: |
133 |
"!!r::('a * 'a)set. single_valued r ==> single_valued (r^n)" |
|
134 |
apply (rule single_valuedI) |
|
135 |
apply (induct n) |
|
136 |
apply simp |
|
137 |
apply (fast dest: single_valuedD elim: rel_pow_Suc_E) |
|
138 |
done |
|
15410 | 139 |
|
140 |
ML |
|
141 |
{* |
|
142 |
val funpow_add = thm "funpow_add"; |
|
143 |
val rel_pow_1 = thm "rel_pow_1"; |
|
144 |
val rel_pow_0_I = thm "rel_pow_0_I"; |
|
145 |
val rel_pow_Suc_I = thm "rel_pow_Suc_I"; |
|
146 |
val rel_pow_Suc_I2 = thm "rel_pow_Suc_I2"; |
|
147 |
val rel_pow_0_E = thm "rel_pow_0_E"; |
|
148 |
val rel_pow_Suc_E = thm "rel_pow_Suc_E"; |
|
149 |
val rel_pow_E = thm "rel_pow_E"; |
|
150 |
val rel_pow_Suc_D2 = thm "rel_pow_Suc_D2"; |
|
151 |
val rel_pow_Suc_D2 = thm "rel_pow_Suc_D2"; |
|
152 |
val rel_pow_E2 = thm "rel_pow_E2"; |
|
153 |
val rtrancl_imp_UN_rel_pow = thm "rtrancl_imp_UN_rel_pow"; |
|
154 |
val rel_pow_imp_rtrancl = thm "rel_pow_imp_rtrancl"; |
|
155 |
val rtrancl_is_UN_rel_pow = thm "rtrancl_is_UN_rel_pow"; |
|
156 |
val single_valued_rel_pow = thm "single_valued_rel_pow"; |
|
157 |
*} |
|
158 |
||
10213 | 159 |
end |