| 3981 |      1 | (*  Title:      HOL/Map.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow, based on a theory by David von Oheimb
 | 
|  |      4 |     Copyright   1997 TU Muenchen
 | 
|  |      5 | 
 | 
|  |      6 | The datatype of `maps' (written ~=>); strongly resembles maps in VDM.
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | Map = List + Option +
 | 
|  |     10 | 
 | 
|  |     11 | types ('a,'b) "~=>" = 'a => 'b option (infixr 0)
 | 
|  |     12 | 
 | 
|  |     13 | consts
 | 
| 5300 |     14 | empty	::  "'a ~=> 'b"
 | 
|  |     15 | chg_map	:: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)"
 | 
| 3981 |     16 | override:: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100)
 | 
| 5300 |     17 | dom	:: "('a ~=> 'b) => 'a set"
 | 
|  |     18 | ran	:: "('a ~=> 'b) => 'b set"
 | 
|  |     19 | map_of	:: "('a * 'b)list => 'a ~=> 'b"
 | 
|  |     20 | map_upds:: "('a ~=> 'b) => 'a list => 'b list => 
 | 
|  |     21 | 	    ('a ~=> 'b)"			 ("_/'(_[|->]_/')" [900,0,0]900)
 | 
|  |     22 | 
 | 
|  |     23 | 
 | 
|  |     24 | syntax
 | 
|  |     25 | map_upd	:: "('a ~=> 'b) => 'a => 'b => ('a ~=> 'b)"
 | 
|  |     26 | 					         ("_/'(_/|->_')"   [900,0,0]900)
 | 
| 3981 |     27 | 
 | 
|  |     28 | syntax (symbols)
 | 
| 5300 |     29 |   "~=>"     :: [type, type] => type      (infixr "\\<leadsto>" 0)
 | 
|  |     30 |   map_upd   :: "('a ~=> 'b) => 'a      => 'b      => ('a ~=> 'b)"
 | 
|  |     31 | 					  ("_/'(_/\\<mapsto>/_')"  [900,0,0]900)
 | 
|  |     32 |   map_upds  :: "('a ~=> 'b) => 'a list => 'b list => ('a ~=> 'b)"
 | 
|  |     33 | 				         ("_/'(_/[\\<mapsto>]/_')" [900,0,0]900)
 | 
| 3981 |     34 |   override  :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)"
 | 
| 5300 |     35 | 				         (infixl "\\<oplus>" 100)
 | 
|  |     36 | 
 | 
|  |     37 | translations
 | 
|  |     38 | 
 | 
|  |     39 |   "m(a|->b)" == "m(a:=Some b)"
 | 
| 3981 |     40 | 
 | 
|  |     41 | defs
 | 
|  |     42 | 
 | 
| 5300 |     43 | empty_def    "empty == %x. None"
 | 
|  |     44 | 
 | 
|  |     45 | chg_map_def  "chg_map f a m == case m a of None => m | Some b => m(a|->f b)"
 | 
| 3981 |     46 | 
 | 
|  |     47 | override_def "m1++m2 == %x. case m2 x of None => m1 x | Some y => Some y"
 | 
|  |     48 | 
 | 
|  |     49 | dom_def "dom(m) == {a. m a ~= None}"
 | 
|  |     50 | ran_def "ran(m) == {b. ? a. m a = Some b}"
 | 
|  |     51 | 
 | 
| 5183 |     52 | primrec
 | 
|  |     53 |   "map_of [] = empty"
 | 
| 5300 |     54 |   "map_of (p#ps) = (map_of ps)(fst p |-> snd p)"
 | 
|  |     55 | 
 | 
|  |     56 | primrec "t([]  [|->]bs) = t"
 | 
|  |     57 |         "t(a#as[|->]bs) = t(a|->hd bs)(as[|->]tl bs)"
 | 
| 3981 |     58 | 
 | 
|  |     59 | end
 |