src/HOL/ex/BinEx.thy
author wenzelm
Tue, 05 Oct 1999 18:26:34 +0200
changeset 7742 01386eb4eab0
parent 6920 c912740c3545
child 9297 bafe45732b10
permissions -rw-r--r--
fixed title;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5545
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     1
(*  Title:      HOL/ex/BinEx.thy
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     2
    ID:         $Id$
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     4
    Copyright   1998  University of Cambridge
5199
be986f7a6def added ex/MonoidGroups (record example);
wenzelm
parents:
diff changeset
     5
5545
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     6
Definition of normal form for proving that binary arithmetic on
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     7
ormalized operands yields normalized results.
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     8
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
     9
Normal means no leading 0s on positive numbers and no leading 1s on negatives.
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    10
*)
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    11
6920
c912740c3545 Introduction of integer division algorithm
paulson
parents: 5545
diff changeset
    12
BinEx = Main +
5545
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    13
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    14
consts normal :: bin set
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    15
  
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    16
inductive "normal"
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    17
  intrs 
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    18
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    19
    Pls  "Pls: normal"
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    20
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    21
    Min  "Min: normal"
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    22
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    23
    BIT_F  "[| w: normal; w ~= Pls |] ==> w BIT False : normal"
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    24
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    25
    BIT_T  "[| w: normal; w ~= Min |] ==> w BIT True : normal"
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    26
9117a0e2bf31 added correctness proofs for arithmetic
paulson
parents: 5199
diff changeset
    27
end