src/HOL/HOLCF/IOA/ABP/Lemmas.thy
author wenzelm
Wed, 04 Oct 2017 12:00:53 +0200
changeset 66787 64b47495676d
parent 62002 f1599e98c4d0
child 67613 ce654b0e6d69
permissions -rw-r--r--
obsolete;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 40945
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/ABP/Lemmas.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     5
theory Lemmas
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
imports Main
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     7
begin
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 42151
diff changeset
     9
subsection \<open>Logic\<close>
19738
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    10
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    11
lemma and_de_morgan_and_absorbe: "(~(A&B)) = ((~A)&B| ~B)"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    12
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    13
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    14
lemma bool_if_impl_or: "(if C then A else B) --> (A|B)"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    15
  by auto
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    16
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    17
lemma exis_elim: "(? x. x=P & Q(x)) = Q(P)"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    18
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    19
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    20
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 42151
diff changeset
    21
subsection \<open>Sets\<close>
19738
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    22
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    23
lemma set_lemmas:
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    24
    "f(x) : (UN x. {f(x)})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    25
    "f x y : (UN x y. {f x y})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    26
    "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    27
    "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    28
  by auto
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    29
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 42151
diff changeset
    30
text \<open>2 Lemmas to add to \<open>set_lemmas\<close>, used also for action handling, 
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 42151
diff changeset
    31
   namely for Intersections and the empty list (compatibility of IOA!).\<close>
19738
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    32
lemma singleton_set: "(UN b.{x. x=f(b)})= (UN b.{f(b)})"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    33
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    34
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    35
lemma de_morgan: "((A|B)=False) = ((~A)&(~B))"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    36
  by blast
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    37
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    38
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 42151
diff changeset
    39
subsection \<open>Lists\<close>
19738
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    40
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    41
lemma cons_not_nil: "l ~= [] --> (? x xs. l = (x#xs))"
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    42
  by (induct l) simp_all
1ac610922636 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    43
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    44
end