src/HOL/IOA/ABP/Lemmas.ML
author oheimb
Fri, 31 May 1996 19:47:23 +0200
changeset 1778 6c942cf3bf68
parent 1465 5d7a7e439cec
child 1894 c2c8279d40f0
permissions -rw-r--r--
adapted some proofs for new simplifier
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1139
993e475e70e2 *** empty log message ***
mueller
parents: 1050
diff changeset
     1
(*  Title:      HOL/IOA/example/Lemmas.ML
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1139
993e475e70e2 *** empty log message ***
mueller
parents: 1050
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
993e475e70e2 *** empty log message ***
mueller
parents: 1050
diff changeset
     4
    Copyright   1994  TU Muenchen
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     5
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     6
*)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     7
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     8
(* Logic *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     9
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    10
val prems = goal HOL.thy "(P ==> Q-->R) ==> P&Q --> R";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    11
  by(fast_tac (HOL_cs addDs prems) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    12
qed "imp_conj_lemma";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    13
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    14
goal HOL.thy "(~(A&B)) = ((~A)&B| ~B)";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    15
by (fast_tac HOL_cs 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    16
val and_de_morgan_and_absorbe = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    17
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    18
goal HOL.thy "(if C then A else B) --> (A|B)";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    19
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    20
by (fast_tac HOL_cs 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    21
val bool_if_impl_or = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    22
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    23
(* Sets *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    24
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    25
val set_lemmas =
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    26
   map (fn s => prove_goal Set.thy s (fn _ => [fast_tac set_cs 1]))
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    27
        ["f(x) : (UN x. {f(x)})",
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    28
         "f x y : (UN x y. {f x y})",
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    29
         "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})",
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    30
         "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"];
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    31
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    32
(* 2 Lemmas to add to set_lemmas ... , used also for action handling, 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    33
   namely for Intersections and the empty list (compatibility of IOA!)  *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    34
goal Set.thy "(UN b.{x.x=f(b)})= (UN b.{f(b)})"; 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    35
 by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    36
 by (fast_tac set_cs 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    37
val singleton_set =result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    38
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    39
goal HOL.thy "((A|B)=False) = ((~A)&(~B))";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    40
 by (fast_tac HOL_cs 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    41
val de_morgan = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    42
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    43
(* Lists *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    44
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1139
diff changeset
    45
val list_ss = simpset_of "List";
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1139
diff changeset
    46
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    47
goal List.thy "hd(l@m) = (if l~=[] then hd(l) else hd(m))";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    48
by (List.list.induct_tac "l" 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    49
by (simp_tac list_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    50
by (simp_tac list_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    51
val hd_append =result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    52
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    53
goal List.thy "l ~= [] --> (? x xs. l = (x#xs))";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    54
 by (List.list.induct_tac "l" 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    55
 by (simp_tac list_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    56
 by (fast_tac HOL_cs 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    57
qed"cons_not_nil";