src/ZF/UNITY/Distributor.thy
author paulson
Thu, 10 Jul 2003 17:14:41 +0200
changeset 14095 a1ba833d6b61
parent 14077 37c964462747
child 16417 9bc16273c2d4
permissions -rw-r--r--
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new variable
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     1
(*  Title: ZF/UNITY/Distributor
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     2
    ID:         $Id$
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     4
    Copyright   2002  University of Cambridge
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     5
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     6
A multiple-client allocator from a single-client allocator:
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     7
Distributor specification
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
     8
*)
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
     9
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    10
theory Distributor = AllocBase + Follows +  Guar + GenPrefix:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    11
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    12
text{*Distributor specification (the number of outputs is Nclients)*}
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    13
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    14
text{*spec (14)*}
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    15
constdefs
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    16
  distr_follows :: "[i, i, i, i =>i] =>i"
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    17
    "distr_follows(A, In, iIn, Out) ==
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    18
     (lift(In) IncreasingWrt prefix(A)/list(A)) \<inter>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    19
     (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \<inter>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    20
     Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients})
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    21
         guarantees
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    22
         (\<Inter>n \<in> Nclients.
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    23
          lift(Out(n))
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    24
              Fols
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    25
          (%s. sublist(s`In, {k \<in> nat. k<length(s`iIn) & nth(k, s`iIn) = n}))
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    26
	  Wrt prefix(A)/list(A))"
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    27
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    28
  distr_allowed_acts :: "[i=>i]=>i"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    29
    "distr_allowed_acts(Out) ==
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    30
     {D \<in> program. AllowedActs(D) =
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    31
     cons(id(state), \<Union>G \<in> (\<Inter>n\<in>nat. preserves(lift(Out(n)))). Acts(G))}"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    32
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    33
  distr_spec :: "[i, i, i, i =>i]=>i"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    34
    "distr_spec(A, In, iIn, Out) ==
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    35
     distr_follows(A, In, iIn, Out) \<inter> distr_allowed_acts(Out)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    36
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    37
locale distr =
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    38
  fixes In  --{*items to distribute*}
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    39
    and iIn --{*destinations of items to distribute*}
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    40
    and Out --{*distributed items*}
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    41
    and A   --{*the type of items being distributed *}
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    42
    and D
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    43
 assumes
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    44
     var_assumes [simp]:  "In \<in> var & iIn \<in> var & (\<forall>n. Out(n):var)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    45
 and all_distinct_vars:   "\<forall>n. all_distinct([In, iIn, iOut(n)])"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    46
 and type_assumes [simp]: "type_of(In)=list(A) &  type_of(iIn)=list(nat) &
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    47
                          (\<forall>n. type_of(Out(n))=list(A))"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    48
 and default_val_assumes [simp]:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    49
                         "default_val(In)=Nil &
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    50
                          default_val(iIn)=Nil &
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    51
                          (\<forall>n. default_val(Out(n))=Nil)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    52
 and distr_spec:  "D \<in> distr_spec(A, In, iIn, Out)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    53
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    54
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    55
lemma (in distr) In_value_type [simp,TC]: "s \<in> state ==> s`In \<in> list(A)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    56
apply (unfold state_def)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    57
apply (drule_tac a = In in apply_type, auto)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    58
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    59
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    60
lemma (in distr) iIn_value_type [simp,TC]:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    61
     "s \<in> state ==> s`iIn \<in> list(nat)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    62
apply (unfold state_def)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    63
apply (drule_tac a = iIn in apply_type, auto)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    64
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    65
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    66
lemma (in distr) Out_value_type [simp,TC]:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    67
     "s \<in> state ==> s`Out(n):list(A)"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    68
apply (unfold state_def)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    69
apply (drule_tac a = "Out (n)" in apply_type)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    70
apply auto
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    71
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    72
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    73
lemma (in distr) D_in_program [simp,TC]: "D \<in> program"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    74
apply (cut_tac distr_spec)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    75
apply (auto simp add: distr_spec_def distr_allowed_acts_def)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    76
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    77
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    78
lemma (in distr) D_ok_iff:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    79
     "G \<in> program ==>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    80
	D ok G <-> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    81
apply (cut_tac distr_spec)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    82
apply (auto simp add: INT_iff distr_spec_def distr_allowed_acts_def
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    83
                      Allowed_def ok_iff_Allowed)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    84
apply (drule safety_prop_Acts_iff [THEN [2] rev_iffD1])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    85
apply (auto intro: safety_prop_Inter)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    86
done
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
    87
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    88
lemma (in distr) distr_Increasing_Out:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    89
"D \<in> ((lift(In) IncreasingWrt prefix(A)/list(A)) \<inter>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    90
      (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \<inter>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    91
      Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iIn). elt<Nclients}))
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    92
      guarantees
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    93
          (\<Inter>n \<in> Nclients. lift(Out(n)) IncreasingWrt
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    94
                            prefix(A)/list(A))"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    95
apply (cut_tac D_in_program distr_spec)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    96
apply (simp (no_asm_use) add: distr_spec_def distr_follows_def)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    97
apply (auto intro!: guaranteesI intro: Follows_imp_Increasing_left 
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    98
            dest!: guaranteesD)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
    99
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   100
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   101
lemma (in distr) distr_bag_Follows_lemma:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   102
"[| \<forall>n \<in> nat. G \<in> preserves(lift(Out(n)));
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   103
   D \<squnion> G \<in> Always({s \<in> state.
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   104
          \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients}) |]
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   105
  ==> D \<squnion> G \<in> Always
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   106
          ({s \<in> state. msetsum(%n. bag_of (sublist(s`In,
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   107
                       {k \<in> nat. k < length(s`iIn) &
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   108
                          nth(k, s`iIn)= n})),
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   109
                         Nclients, A) =
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   110
              bag_of(sublist(s`In, length(s`iIn)))})"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   111
apply (cut_tac D_in_program)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   112
apply (subgoal_tac "G \<in> program")
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   113
 prefer 2 apply (blast dest: preserves_type [THEN subsetD])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   114
apply (erule Always_Diff_Un_eq [THEN iffD1])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   115
apply (rule state_AlwaysI [THEN Always_weaken], auto)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   116
apply (rename_tac s)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   117
apply (subst bag_of_sublist_UN_disjoint [symmetric])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   118
   apply (simp (no_asm_simp) add: nat_into_Finite)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   119
  apply blast
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   120
 apply (simp (no_asm_simp))
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   121
apply (simp add: set_of_list_conv_nth [of _ nat])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   122
apply (subgoal_tac
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   123
       "(\<Union>i \<in> Nclients. {k\<in>nat. k < length(s`iIn) & nth(k, s`iIn) = i}) =
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   124
        length(s`iIn) ")
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   125
apply (simp (no_asm_simp))
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   126
apply (rule equalityI)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   127
apply (force simp add: ltD, safe)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   128
apply (rename_tac m)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   129
apply (subgoal_tac "length (s ` iIn) \<in> nat")
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   130
apply typecheck
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   131
apply (subgoal_tac "m \<in> nat")
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   132
apply (drule_tac x = "nth(m, s`iIn) " and P = "%elt. ?X (elt) --> elt<Nclients" in bspec)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   133
apply (simp add: ltI)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   134
apply (simp_all add: Ord_mem_iff_lt)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   135
apply (blast dest: ltD)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   136
apply (blast intro: lt_trans)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   137
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   138
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   139
theorem (in distr) distr_bag_Follows:
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   140
 "D \<in> ((lift(In) IncreasingWrt prefix(A)/list(A)) \<inter>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   141
       (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \<inter>
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   142
        Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients}))
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   143
      guarantees
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   144
       (\<Inter>n \<in> Nclients.
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   145
        (%s. msetsum(%i. bag_of(s`Out(i)),  Nclients, A))
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   146
        Fols
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   147
        (%s. bag_of(sublist(s`In, length(s`iIn))))
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   148
        Wrt MultLe(A, r)/Mult(A))"
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   149
apply (cut_tac distr_spec)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   150
apply (rule guaranteesI, clarify)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   151
apply (rule distr_bag_Follows_lemma [THEN Always_Follows2])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   152
apply (simp add: D_ok_iff, auto)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   153
apply (rule Follows_state_ofD1)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   154
apply (rule Follows_msetsum_UN)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   155
   apply (simp_all add: nat_into_Finite bag_of_multiset [of _ A])
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   156
apply (auto simp add: distr_spec_def distr_follows_def)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   157
apply (drule guaranteesD, assumption)
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   158
apply (simp_all cong add: Follows_cong
14077
37c964462747 Conversion of theory UNITY to Isar script
paulson
parents: 14072
diff changeset
   159
		add: refl_prefix
37c964462747 Conversion of theory UNITY to Isar script
paulson
parents: 14072
diff changeset
   160
		   mono_bag_of [THEN subset_Follows_comp, THEN subsetD, 
37c964462747 Conversion of theory UNITY to Isar script
paulson
parents: 14072
diff changeset
   161
				unfolded metacomp_def])
14072
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   162
done
f932be305381 Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents: 14057
diff changeset
   163
14052
e9c9f69e4f63 some new ZF/UNITY material from Sidi Ehmety
paulson
parents:
diff changeset
   164
end