author | wenzelm |
Thu, 22 Nov 2007 14:51:34 +0100 | |
changeset 25456 | 6f79698f294d |
parent 19738 | 1ac610922636 |
child 26649 | a053f13bc9da |
permissions | -rw-r--r-- |
3072
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
1 |
(* Title: HOLCF/IOA/ABP/Lemmas.thy |
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
2 |
ID: $Id$ |
12218 | 3 |
Author: Olaf Müller |
3072
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
4 |
*) |
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
5 |
|
17244 | 6 |
theory Lemmas |
7 |
imports Main |
|
8 |
begin |
|
9 |
||
19738 | 10 |
subsection {* Logic *} |
11 |
||
12 |
lemma and_de_morgan_and_absorbe: "(~(A&B)) = ((~A)&B| ~B)" |
|
13 |
by blast |
|
14 |
||
15 |
lemma bool_if_impl_or: "(if C then A else B) --> (A|B)" |
|
16 |
by auto |
|
17 |
||
18 |
lemma exis_elim: "(? x. x=P & Q(x)) = Q(P)" |
|
19 |
by blast |
|
20 |
||
21 |
||
22 |
subsection {* Sets *} |
|
23 |
||
24 |
lemma set_lemmas: |
|
25 |
"f(x) : (UN x. {f(x)})" |
|
26 |
"f x y : (UN x y. {f x y})" |
|
27 |
"!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})" |
|
28 |
"!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})" |
|
29 |
by auto |
|
30 |
||
31 |
text {* 2 Lemmas to add to @{text "set_lemmas"}, used also for action handling, |
|
32 |
namely for Intersections and the empty list (compatibility of IOA!). *} |
|
33 |
lemma singleton_set: "(UN b.{x. x=f(b)})= (UN b.{f(b)})" |
|
34 |
by blast |
|
35 |
||
36 |
lemma de_morgan: "((A|B)=False) = ((~A)&(~B))" |
|
37 |
by blast |
|
38 |
||
39 |
||
40 |
subsection {* Lists *} |
|
41 |
||
42 |
lemma hd_append: "hd(l@m) = (if l~=[] then hd(l) else hd(m))" |
|
43 |
by (induct l) simp_all |
|
44 |
||
45 |
lemma cons_not_nil: "l ~= [] --> (? x xs. l = (x#xs))" |
|
46 |
by (induct l) simp_all |
|
47 |
||
17244 | 48 |
end |