src/HOLCF/IOA/ABP/Lemmas.ML
author wenzelm
Tue, 27 Aug 2002 11:03:05 +0200
changeset 13524 604d0f3622d6
parent 12218 6597093b77e7
child 14401 477380c74c1d
permissions -rw-r--r--
*** empty log message ***
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.ML
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 5192
diff changeset
     3
    Author:     Olaf Müller
wenzelm
parents: 5192
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     5
*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     6
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     7
(* Logic *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     8
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     9
goal HOL.thy "(~(A&B)) = ((~A)&B| ~B)";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    10
by (Fast_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
val and_de_morgan_and_absorbe = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    12
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
goal HOL.thy "(if C then A else B) --> (A|B)";
4833
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4423
diff changeset
    14
by (stac split_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    15
by (Fast_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    16
val bool_if_impl_or = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    18
(* Sets *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    19
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    20
val set_lemmas =
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    21
   map (fn s => prove_goal Set.thy s (fn _ => [Fast_tac 1]))
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    22
        ["f(x) : (UN x. {f(x)})",
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    23
         "f x y : (UN x y. {f x y})",
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    24
         "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})",
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    25
         "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    26
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
(* 2 Lemmas to add to set_lemmas ... , used also for action handling, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    28
   namely for Intersections and the empty list (compatibility of IOA!)  *)
3852
e694c660055b fixed dots;
wenzelm
parents: 3072
diff changeset
    29
goal Set.thy "(UN b.{x. x=f(b)})= (UN b.{f(b)})"; 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    30
 by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    31
 by (Fast_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    32
val singleton_set =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    33
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    34
goal HOL.thy "((A|B)=False) = ((~A)&(~B))";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    35
 by (Fast_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    36
val de_morgan = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    37
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    38
(* Lists *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    39
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3852
diff changeset
    40
val list_ss = simpset_of List.thy;
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    41
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    42
goal List.thy "hd(l@m) = (if l~=[] then hd(l) else hd(m))";
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4833
diff changeset
    43
by (induct_tac "l" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    44
by (simp_tac list_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    45
by (simp_tac list_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    46
val hd_append =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    47
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    48
goal List.thy "l ~= [] --> (? x xs. l = (x#xs))";
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4833
diff changeset
    49
 by (induct_tac "l" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    50
 by (simp_tac list_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    51
 by (Fast_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    52
qed"cons_not_nil";