src/HOL/Fun.thy
author wenzelm
Sat, 21 Nov 1998 12:17:18 +0100
changeset 5944 dcc446da8e19
parent 5852 4d7320490be4
child 6171 cd237a10cbf8
permissions -rw-r--r--
added undos, redos;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 923
diff changeset
     1
(*  Title:      HOL/Fun.thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 923
diff changeset
     3
    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
     6
Notions about functions.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
5852
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
     9
Fun = Vimage + equalities + 
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    10
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 2912
diff changeset
    11
instance set :: (term) order
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 2912
diff changeset
    12
                       (subset_refl,subset_trans,subset_antisym,psubset_eq)
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    13
consts
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    14
5608
a82a038a3e7a id <-> Id
nipkow
parents: 5305
diff changeset
    15
  id          ::  'a => 'a
5305
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    16
  o           :: ['b => 'c, 'a => 'b, 'a] => 'c   (infixl 55)
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4648
diff changeset
    17
  inj, surj   :: ('a => 'b) => bool                   (*inj/surjective*)
bd73675adbed Added a few lemmas.
nipkow
parents: 4648
diff changeset
    18
  inj_on      :: ['a => 'b, 'a set] => bool
bd73675adbed Added a few lemmas.
nipkow
parents: 4648
diff changeset
    19
  inv         :: ('a => 'b) => ('b => 'a)
5305
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    20
  fun_upd  :: "('a => 'b) => 'a => 'b => ('a => 'b)"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    21
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    22
nonterminals
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    23
  updbinds  updbind
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    24
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    25
syntax
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    26
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    27
  (* Let expressions *)
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    28
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    29
  "_updbind"       :: ['a, 'a] => updbind             ("(2_ :=/ _)")
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    30
  ""               :: updbind => updbinds             ("_")
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    31
  "_updbinds"      :: [updbind, updbinds] => updbinds ("_,/ _")
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    32
  "_Update"        :: ['a, updbinds] => 'a            ("_/'((_)')" [900,0] 900)
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    33
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    34
translations
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    35
  "_Update f (_updbinds b bs)"  == "_Update (_Update f b) bs"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    36
  "f(x:=y)"                     == "fun_upd f x y"
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    37
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    38
defs
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    39
5608
a82a038a3e7a id <-> Id
nipkow
parents: 5305
diff changeset
    40
  id_def	"id             == %x. x"
5305
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    41
  o_def   	"f o g          == %x. f(g(x))"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    42
  inj_def	"inj f          == ! x y. f(x)=f(y) --> x=y"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    43
  inj_on_def	"inj_on f A     == ! x:A. ! y:A. f(x)=f(y) --> x=y"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    44
  surj_def	"surj f         == ! y. ? x. y=f(x)"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    45
  inv_def	"inv(f::'a=>'b) == % y. @x. f(x)=y"
513925de8962 cleanup for Fun.thy:
oheimb
parents: 4830
diff changeset
    46
  fun_upd_def	"f(a:=b)        == % x. if x=a then b else f x"
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    47
5852
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    48
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    49
(*The Pi-operator, by Florian Kammueller*)
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    50
  
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    51
constdefs
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    52
  Pi      :: "['a set, 'a => 'b set] => ('a => 'b) set"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    53
    "Pi A B == {f. ! x. if x:A then f(x) : B(x) else f(x) = (@ y. True)}"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    54
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    55
  restrict :: "['a => 'b, 'a set] => ('a => 'b)"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    56
    "restrict f A == (%x. if x : A then f x else (@ y. True))"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    57
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    58
syntax
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    59
  "@Pi"  :: "[idt, 'a set, 'b set] => ('a => 'b) set"  ("(3PI _:_./ _)" 10)
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    60
  funcset :: "['a set, 'b set] => ('a => 'b) set"      (infixr 60) 
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    61
  "@lam" :: "[pttrn, 'a set, 'a => 'b] => ('a => 'b)"  ("(3lam _:_./ _)" 10)
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    62
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    63
  (*Giving funcset the nice arrow syntax -> clashes with existing theories*)
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    64
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    65
translations
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    66
  "PI x:A. B" => "Pi A (%x. B)"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    67
  "A funcset B"    => "Pi A (_K B)"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    68
  "lam x:A. f"  == "restrict (%x. f) A"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    69
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    70
constdefs
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    71
  Applyall :: "[('a => 'b) set, 'a]=> 'b set"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    72
    "Applyall F a == (%f. f a) `` F"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    73
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    74
  compose :: "['a set, 'a => 'b, 'b => 'c] => ('a => 'c)"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    75
    "compose A g f == lam x : A. g(f x)"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    76
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    77
  Inv    :: "['a set, 'a => 'b] => ('b => 'a)"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    78
    "Inv A f == (% x. (@ y. y : A & f y = x))"
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    79
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    80
  
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 1475
diff changeset
    81
end
5852
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    82
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    83
ML
4d7320490be4 the function space operator
paulson
parents: 5608
diff changeset
    84
val print_translation = [("Pi", dependent_tr' ("@Pi", "op funcset"))];