src/HOLCF/IOA/NTP/Multiset.thy
author wenzelm
Mon, 23 Jun 2008 16:01:03 +0200
changeset 27327 efd626efcb04
parent 25161 aa8474398030
child 35174 e15040ae75d7
permissions -rw-r--r--
info: default name is "", not "Pure";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
theory Multiset
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
imports Lemmas
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    10
begin
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    17
  "{|}"  :: "'a multiset"                        ("{|}")
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    18
  addm   :: "['a multiset, 'a] => 'a multiset"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    19
  delm   :: "['a multiset, 'a] => 'a multiset"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    20
  countm :: "['a multiset, 'a => bool] => nat"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    23
axioms
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    24
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    25
delm_empty_def:
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    26
  "delm {|} x = {|}"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    37
count_def:
3852
e694c660055b fixed dots;
wenzelm
parents: 3073
diff changeset
    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
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    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
19739
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    43
lemma count_empty: 
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    44
   "count {|} x = 0"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    45
  by (simp add: Multiset.count_def Multiset.countm_empty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    46
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    47
lemma count_addm_simp: 
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    48
     "count (addm M x) y = (if y=x then Suc(count M y) else count M y)"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    49
  by (simp add: Multiset.count_def Multiset.countm_nonempty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    50
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    51
lemma count_leq_addm: "count M y <= count (addm M x) y"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    52
  by (simp add: count_addm_simp)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    53
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    54
lemma count_delm_simp: 
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    55
     "count (delm M x) y = (if y=x then count M y - 1 else count M y)"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    56
apply (unfold Multiset.count_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    57
apply (rule_tac M = "M" in Multiset.induction)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    58
apply (simp (no_asm_simp) add: Multiset.delm_empty_def Multiset.countm_empty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    59
apply (simp add: Multiset.delm_nonempty_def Multiset.countm_nonempty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    60
apply safe
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    61
apply simp
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    62
done
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    63
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    64
lemma countm_props: "!!M. (!x. P(x) --> Q(x)) ==> (countm M P <= countm M Q)"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    65
apply (rule_tac M = "M" in Multiset.induction)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    66
 apply (simp (no_asm) add: Multiset.countm_empty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    67
apply (simp (no_asm) add: Multiset.countm_nonempty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    68
apply auto
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    69
done
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    70
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    71
lemma countm_spurious_delm: "!!P. ~P(obj) ==> countm M P = countm (delm M obj) P"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    72
  apply (rule_tac M = "M" in Multiset.induction)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    73
  apply (simp (no_asm) add: Multiset.delm_empty_def Multiset.countm_empty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    74
  apply (simp (no_asm_simp) add: Multiset.countm_nonempty_def Multiset.delm_nonempty_def)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    75
  done
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    76
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    77
25161
aa8474398030 changed back from ~=0 to >0
nipkow
parents: 25141
diff changeset
    78
lemma pos_count_imp_pos_countm [rule_format (no_asm)]: "!!P. P(x) ==> 0<count M x --> countm M P > 0"
19739
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    79
  apply (rule_tac M = "M" in Multiset.induction)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    80
  apply (simp (no_asm) add: Multiset.delm_empty_def Multiset.count_def Multiset.countm_empty_def)
25161
aa8474398030 changed back from ~=0 to >0
nipkow
parents: 25141
diff changeset
    81
  apply (simp add: Multiset.count_def Multiset.delm_nonempty_def Multiset.countm_nonempty_def)
19739
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    82
  done
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    83
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    84
lemma countm_done_delm: 
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    85
   "!!P. P(x) ==> 0<count M x --> countm (delm M x) P = countm M P - 1"
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    86
  apply (rule_tac M = "M" in Multiset.induction)
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    87
  apply (simp (no_asm) add: Multiset.delm_empty_def Multiset.countm_empty_def)
25161
aa8474398030 changed back from ~=0 to >0
nipkow
parents: 25141
diff changeset
    88
  apply (simp (no_asm_simp) add: count_addm_simp Multiset.delm_nonempty_def Multiset.countm_nonempty_def pos_count_imp_pos_countm)
19739
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    89
  apply auto
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    90
  done
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    91
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    92
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    93
declare count_addm_simp [simp] count_delm_simp [simp]
c58ef2aa5430 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    94
  Multiset.countm_empty_def [simp] Multiset.delm_empty_def [simp] count_empty [simp]
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    95
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    96
end