author | wenzelm |
Wed, 04 Oct 2017 12:00:53 +0200 | |
changeset 66787 | 64b47495676d |
parent 62002 | f1599e98c4d0 |
child 67613 | ce654b0e6d69 |
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 |
||
62002 | 9 |
subsection \<open>Logic\<close> |
19738 | 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 |
||
62002 | 21 |
subsection \<open>Sets\<close> |
19738 | 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 |
||
62002 | 30 |
text \<open>2 Lemmas to add to \<open>set_lemmas\<close>, used also for action handling, |
31 |
namely for Intersections and the empty list (compatibility of IOA!).\<close> |
|
19738 | 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 |
||
62002 | 39 |
subsection \<open>Lists\<close> |
19738 | 40 |
|
41 |
lemma cons_not_nil: "l ~= [] --> (? x xs. l = (x#xs))" |
|
42 |
by (induct l) simp_all |
|
43 |
||
17244 | 44 |
end |