src/ZF/ex/BinEx.ML
author wenzelm
Sun, 25 Oct 1998 12:33:27 +0100
changeset 5769 6a422b22ba02
parent 5533 bce36a019b03
child 6153 bff90585cce5
permissions -rw-r--r--
tuned checklist;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5533
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ex/BinEx.ML
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     2
    ID:         $Id$
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     5
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     6
Examples of performing binary arithmetic by simplification
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     7
*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     8
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
     9
context Bin.thy;
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    10
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    11
set proof_timing;
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    12
(*All runtimes below are on a 300MHz Pentium*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    13
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    14
Goal "#13  $+  #19 = #32";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    15
by (simp_tac bin_comp_ss 1);    (*0 secs*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    16
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    17
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    18
Goal "#1234  $+  #5678 = #6912";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    19
by (simp_tac bin_comp_ss 1);    (*190 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    20
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    21
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    22
Goal "#1359  $+  #-2468 = #-1109";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    23
by (simp_tac bin_comp_ss 1);    (*160 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    24
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    25
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    26
Goal "#93746  $+  #-46375 = #47371";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    27
by (simp_tac bin_comp_ss 1);    (*300 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    28
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    29
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    30
Goal "$~ #65745 = #-65745";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    31
by (simp_tac bin_comp_ss 1);    (*80 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    32
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    33
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    34
(* negation of ~54321 *)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    35
Goal "$~ #-54321 = #54321";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    36
by (simp_tac bin_comp_ss 1);    (*90 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    37
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    38
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    39
Goal "#13  $*  #19 = #247";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    40
by (simp_tac bin_comp_ss 1);    (*110 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    41
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    42
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    43
Goal "#-84  $*  #51 = #-4284";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    44
by (simp_tac bin_comp_ss 1);    (*210 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    45
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    46
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    47
(*The worst case for 8-bit operands *)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    48
Goal "#255  $*  #255 = #65025";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    49
by (simp_tac bin_comp_ss 1);    (*730 msec*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    50
result();
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    51
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    52
Goal "#1359  $*  #-2468 = #-3354012";
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    53
by (simp_tac bin_comp_ss 1);    (*1.04 secs*)
bce36a019b03 re-organized for the new directory Integ
paulson
parents:
diff changeset
    54
result();