| author | wenzelm | 
| Sat, 07 Jan 2006 23:27:56 +0100 | |
| changeset 18616 | cf5d07758d3f | 
| parent 17244 | 0b2ff9541727 | 
| child 19739 | c58ef2aa5430 | 
| permissions | -rw-r--r-- | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 1 | (* Title: HOL/IOA/NTP/Multiset.thy | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 2 | ID: $Id$ | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | Author: Tobias Nipkow & Konrad Slind | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | *) | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 5 | |
| 17244 | 6 | header {* Axiomatic multisets *}
 | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 7 | |
| 17244 | 8 | theory Multiset | 
| 9 | imports Lemmas | |
| 10 | begin | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 11 | |
| 17244 | 12 | typedecl | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 13 | 'a multiset | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 14 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 15 | consts | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 16 | |
| 17244 | 17 |   "{|}"  :: "'a multiset"                        ("{|}")
 | 
| 18 | addm :: "['a multiset, 'a] => 'a multiset" | |
| 19 | delm :: "['a multiset, 'a] => 'a multiset" | |
| 20 | countm :: "['a multiset, 'a => bool] => nat" | |
| 21 | count :: "['a multiset, 'a] => nat" | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 22 | |
| 17244 | 23 | axioms | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 24 | |
| 17244 | 25 | delm_empty_def: | 
| 26 |   "delm {|} x = {|}"
 | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 27 | |
| 17244 | 28 | delm_nonempty_def: | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 29 | "delm (addm M x) y == (if x=y then M else addm (delm M y) x)" | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 30 | |
| 17244 | 31 | countm_empty_def: | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 32 |    "countm {|} P == 0"
 | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 33 | |
| 17244 | 34 | countm_nonempty_def: | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 35 | "countm (addm M x) P == countm M P + (if P x then Suc 0 else 0)" | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 36 | |
| 17244 | 37 | count_def: | 
| 3852 | 38 | "count M x == countm M (%y. y = x)" | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 39 | |
| 17244 | 40 | "induction": | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 41 |    "[| P({|}); !!M x. P(M) ==> P(addm M x) |] ==> P(M)"
 | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 42 | |
| 17244 | 43 | ML {* use_text Context.ml_output true
 | 
| 44 |   ("structure Multiset = struct " ^ legacy_bindings (the_context ()) ^ " end") *}
 | |
| 45 | ML {* open Multiset *}
 | |
| 46 | ||
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 47 | end |