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