src/HOL/UNITY/LessThan.thy
author oheimb
Fri, 01 May 1998 22:30:42 +0200
changeset 4884 1ec740e30811
parent 4776 1f9362e769c1
child 5232 e5a7cdd07ea5
permissions -rw-r--r--
Auto_tac: now uses enhanced version of asm_full_simp_tac, Force_tac: replaced fast_tac by best_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/LessThan
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
lessThan, greaterThan, atLeast, atMost
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
LessThan = List +
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
constdefs
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
  lessThan   :: "nat => nat set"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
     "lessThan n == {i. i<n}"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
  atMost   :: "nat => nat set"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
     "atMost n == {i. i<=n}"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
  greaterThan   :: "nat => nat set"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
     "greaterThan n == {i. n<i}"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
  atLeast   :: "nat => nat set"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
     "atLeast n == {i. n<=i}"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
end