src/HOLCF/IOA/ABP/Lemmas.thy
author wenzelm
Thu, 22 Nov 2007 14:51:34 +0100
changeset 25456 6f79698f294d
parent 19738 1ac610922636
child 26649 a053f13bc9da
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
wenzelm
parents: 10215
diff changeset
     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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
theory Lemmas
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     7
imports Main
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
begin
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
19738
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    10
subsection {* Logic *}
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    11
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    12
lemma and_de_morgan_and_absorbe: "(~(A&B)) = ((~A)&B| ~B)"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    13
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    14
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    15
lemma bool_if_impl_or: "(if C then A else B) --> (A|B)"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    16
  by auto
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    17
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    18
lemma exis_elim: "(? x. x=P & Q(x)) = Q(P)"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    19
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    20
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    21
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    22
subsection {* Sets *}
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    23
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    24
lemma set_lemmas:
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    25
    "f(x) : (UN x. {f(x)})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    26
    "f x y : (UN x y. {f x y})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    27
    "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    28
    "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    29
  by auto
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    30
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    31
text {* 2 Lemmas to add to @{text "set_lemmas"}, used also for action handling, 
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    32
   namely for Intersections and the empty list (compatibility of IOA!). *}
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    33
lemma singleton_set: "(UN b.{x. x=f(b)})= (UN b.{f(b)})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    34
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    35
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    36
lemma de_morgan: "((A|B)=False) = ((~A)&(~B))"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    37
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    38
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    39
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    40
subsection {* Lists *}
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    41
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    42
lemma hd_append: "hd(l@m) = (if l~=[] then hd(l) else hd(m))"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    43
  by (induct l) simp_all
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    44
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    45
lemma cons_not_nil: "l ~= [] --> (? x xs. l = (x#xs))"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    46
  by (induct l) simp_all
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    47
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    48
end