src/ZF/UNITY/WFair.thy
author paulson
Mon, 28 Mar 2005 16:19:56 +0200
changeset 15634 bca33c49b083
parent 14077 37c964462747
child 24893 b8ef7afe3a6b
permissions -rw-r--r--
conversion of UNITY to Isar scripts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/WFair.thy
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
     7
header{*Progress under Weak Fairness*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
     8
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
     9
theory WFair
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    10
imports UNITY Main_ZFC
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    11
begin
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    12
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    13
text{*This theory defines the operators transient, ensures and leadsTo,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    14
assuming weak fairness. From Misra, "A Logic for Concurrent Programming",
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    15
1994.*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    16
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
  
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    19
  (* This definition specifies weak fairness.  The rest of the theory
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
    is generic to all forms of fairness.*) 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
 transient :: "i=>i"
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    22
  "transient(A) =={F:program. (EX act: Acts(F). A<=domain(act) &
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    23
			       act``A <= state-A) & st_set(A)}"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    25
  ensures :: "[i,i] => i"       (infixl "ensures" 60)
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
  "A ensures B == ((A-B) co (A Un B)) Int transient(A-B)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
consts
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
  (*LEADS-TO constant for the inductive definition*)
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    31
  leads :: "[i, i]=>i"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    33
inductive 
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
  domains
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    35
     "leads(D, F)" <= "Pow(D)*Pow(D)"
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    36
  intros 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    37
    Basis:  "[| F:A ensures B;  A:Pow(D); B:Pow(D) |] ==> <A,B>:leads(D, F)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    38
    Trans:  "[| <A,B> : leads(D, F); <B,C> : leads(D, F) |] ==>  <A,C>:leads(D, F)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    39
    Union:   "[| S:Pow({A:S. <A, B>:leads(D, F)}); B:Pow(D); S:Pow(Pow(D)) |] ==> 
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    40
	      <Union(S),B>:leads(D, F)"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
  monos        Pow_mono
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    43
  type_intros  Union_Pow_iff [THEN iffD2] UnionI PowI
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    47
  (* The Visible version of the LEADS-TO relation*)
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    48
  leadsTo :: "[i, i] => i"       (infixl "leadsTo" 60)
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    49
  "A leadsTo B == {F:program. <A,B>:leads(state, F)}"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
  
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    51
  (* wlt(F, B) is the largest set that leads to B*)
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
  wlt :: "[i, i] => i"
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    53
    "wlt(F, B) == Union({A:Pow(state). F: A leadsTo B})"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
syntax (xsymbols)
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    56
  "leadsTo" :: "[i, i] => i" (infixl "\<longmapsto>" 60)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    57
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    58
(** Ad-hoc set-theory rules **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    59
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    60
lemma Int_Union_Union: "Union(B) Int A = (\<Union>b \<in> B. b Int A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    61
by auto
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    62
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    63
lemma Int_Union_Union2: "A Int Union(B) = (\<Union>b \<in> B. A Int b)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    64
by auto
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    65
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    66
(*** transient ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    67
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    68
lemma transient_type: "transient(A)<=program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    69
by (unfold transient_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    70
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    71
lemma transientD2: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    72
"F \<in> transient(A) ==> F \<in> program & st_set(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    73
apply (unfold transient_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    74
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    75
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    76
lemma stable_transient_empty: "[| F \<in> stable(A); F \<in> transient(A) |] ==> A = 0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    77
by (simp add: stable_def constrains_def transient_def, fast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    78
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    79
lemma transient_strengthen: "[|F \<in> transient(A); B<=A|] ==> F \<in> transient(B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    80
apply (simp add: transient_def st_set_def, clarify)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    81
apply (blast intro!: rev_bexI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    82
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    83
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    84
lemma transientI: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    85
"[|act \<in> Acts(F); A <= domain(act); act``A <= state-A;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    86
    F \<in> program; st_set(A)|] ==> F \<in> transient(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    87
by (simp add: transient_def, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    88
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    89
lemma transientE: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    90
     "[| F \<in> transient(A);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    91
         !!act. [| act \<in> Acts(F);  A <= domain(act); act``A <= state-A|]==>P|]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    92
      ==>P"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    93
by (simp add: transient_def, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    94
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    95
lemma transient_state: "transient(state) = 0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    96
apply (simp add: transient_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    97
apply (rule equalityI, auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    98
apply (cut_tac F = x in Acts_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
    99
apply (simp add: Diff_cancel)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   100
apply (auto intro: st0_in_state)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   101
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   102
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   103
lemma transient_state2: "state<=B ==> transient(B) = 0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   104
apply (simp add: transient_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   105
apply (rule equalityI, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   106
apply (cut_tac F = x in Acts_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   107
apply (subgoal_tac "B=state")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   108
apply (auto intro: st0_in_state)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   109
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   110
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   111
lemma transient_empty: "transient(0) = program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   112
by (auto simp add: transient_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   113
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   114
declare transient_empty [simp] transient_state [simp] transient_state2 [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   115
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   116
(*** ensures ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   117
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   118
lemma ensures_type: "A ensures B <=program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   119
by (simp add: ensures_def constrains_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   120
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   121
lemma ensuresI: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   122
"[|F:(A-B) co (A Un B); F \<in> transient(A-B)|]==>F \<in> A ensures B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   123
apply (unfold ensures_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   124
apply (auto simp add: transient_type [THEN subsetD])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   125
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   126
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   127
(* Added by Sidi, from Misra's notes, Progress chapter, exercise 4 *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   128
lemma ensuresI2: "[| F \<in> A co A Un B; F \<in> transient(A) |] ==> F \<in> A ensures B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   129
apply (drule_tac B = "A-B" in constrains_weaken_L)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   130
apply (drule_tac [2] B = "A-B" in transient_strengthen)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   131
apply (auto simp add: ensures_def transient_type [THEN subsetD])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   132
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   133
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   134
lemma ensuresD: "F \<in> A ensures B ==> F:(A-B) co (A Un B) & F \<in> transient (A-B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   135
by (unfold ensures_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   136
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   137
lemma ensures_weaken_R: "[|F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   138
apply (unfold ensures_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   139
apply (blast intro: transient_strengthen constrains_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   140
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   141
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   142
(*The L-version (precondition strengthening) fails, but we have this*) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   143
lemma stable_ensures_Int: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   144
     "[| F \<in> stable(C);  F \<in> A ensures B |] ==> F:(C Int A) ensures (C Int B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   145
 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   146
apply (unfold ensures_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   147
apply (simp (no_asm) add: Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   148
apply (blast intro: transient_strengthen stable_constrains_Int constrains_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   149
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   150
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   151
lemma stable_transient_ensures: "[|F \<in> stable(A);  F \<in> transient(C); A<=B Un C|] ==> F \<in> A ensures B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   152
apply (frule stable_type [THEN subsetD])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   153
apply (simp add: ensures_def stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   154
apply (blast intro: transient_strengthen constrains_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   155
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   156
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   157
lemma ensures_eq: "(A ensures B) = (A unless B) Int transient (A-B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   158
by (auto simp add: ensures_def unless_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   159
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   160
lemma subset_imp_ensures: "[| F \<in> program; A<=B  |] ==> F \<in> A ensures B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   161
by (auto simp add: ensures_def constrains_def transient_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   162
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   163
(*** leadsTo ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   164
lemmas leads_left = leads.dom_subset [THEN subsetD, THEN SigmaD1]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   165
lemmas leads_right = leads.dom_subset [THEN subsetD, THEN SigmaD2]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   166
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   167
lemma leadsTo_type: "A leadsTo B <= program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   168
by (unfold leadsTo_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   169
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   170
lemma leadsToD2: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   171
"F \<in> A leadsTo B ==> F \<in> program & st_set(A) & st_set(B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   172
apply (unfold leadsTo_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   173
apply (blast dest: leads_left leads_right)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   174
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   175
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   176
lemma leadsTo_Basis: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   177
    "[|F \<in> A ensures B; st_set(A); st_set(B)|] ==> F \<in> A leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   178
apply (unfold leadsTo_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   179
apply (cut_tac ensures_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   180
apply (auto intro: leads.Basis)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   181
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   182
declare leadsTo_Basis [intro]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   183
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   184
(* Added by Sidi, from Misra's notes, Progress chapter, exercise number 4 *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   185
(* [| F \<in> program; A<=B;  st_set(A); st_set(B) |] ==> A leadsTo B *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   186
lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis, standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   187
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   188
lemma leadsTo_Trans: "[|F \<in> A leadsTo B;  F \<in> B leadsTo C |]==>F \<in> A leadsTo C"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   189
apply (unfold leadsTo_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   190
apply (auto intro: leads.Trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   191
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   192
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   193
(* Better when used in association with leadsTo_weaken_R *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   194
lemma transient_imp_leadsTo: "F \<in> transient(A) ==> F \<in> A leadsTo (state-A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   195
apply (unfold transient_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   196
apply (blast intro: ensuresI [THEN leadsTo_Basis] constrains_weaken transientI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   197
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   198
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   199
(*Useful with cancellation, disjunction*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   200
lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' Un A') ==> F \<in> A leadsTo A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   201
by simp
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   202
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   203
lemma leadsTo_Un_duplicate2:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   204
     "F \<in> A leadsTo (A' Un C Un C) ==> F \<in> A leadsTo (A' Un C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   205
by (simp add: Un_ac)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   206
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   207
(*The Union introduction rule as we should have liked to state it*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   208
lemma leadsTo_Union: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   209
    "[|!!A. A \<in> S ==> F \<in> A leadsTo B; F \<in> program; st_set(B)|]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   210
     ==> F \<in> Union(S) leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   211
apply (unfold leadsTo_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   212
apply (blast intro: leads.Union dest: leads_left)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   213
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   214
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   215
lemma leadsTo_Union_Int: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   216
    "[|!!A. A \<in> S ==>F : (A Int C) leadsTo B; F \<in> program; st_set(B)|]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   217
     ==> F : (Union(S)Int C)leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   218
apply (unfold leadsTo_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   219
apply (simp only: Int_Union_Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   220
apply (blast dest: leads_left intro: leads.Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   221
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   222
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   223
lemma leadsTo_UN: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   224
    "[| !!i. i \<in> I ==> F \<in> A(i) leadsTo B; F \<in> program; st_set(B)|]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   225
     ==> F:(\<Union>i \<in> I. A(i)) leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   226
apply (simp add: Int_Union_Union leadsTo_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   227
apply (blast dest: leads_left intro: leads.Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   228
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   229
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   230
(* Binary union introduction rule *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   231
lemma leadsTo_Un:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   232
     "[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A Un B) leadsTo C"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   233
apply (subst Un_eq_Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   234
apply (blast intro: leadsTo_Union dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   235
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   236
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   237
lemma single_leadsTo_I:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   238
    "[|!!x. x \<in> A==> F:{x} leadsTo B; F \<in> program; st_set(B) |] 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   239
     ==> F \<in> A leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   240
apply (rule_tac b = A in UN_singleton [THEN subst])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   241
apply (rule leadsTo_UN, auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   242
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   243
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   244
lemma leadsTo_refl: "[| F \<in> program; st_set(A) |] ==> F \<in> A leadsTo A"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   245
by (blast intro: subset_imp_leadsTo)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   246
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   247
lemma leadsTo_refl_iff: "F \<in> A leadsTo A <-> F \<in> program & st_set(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   248
by (auto intro: leadsTo_refl dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   249
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   250
lemma empty_leadsTo: "F \<in> 0 leadsTo B <-> (F \<in> program & st_set(B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   251
by (auto intro: subset_imp_leadsTo dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   252
declare empty_leadsTo [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   253
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   254
lemma leadsTo_state: "F \<in> A leadsTo state <-> (F \<in> program & st_set(A))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   255
by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   256
declare leadsTo_state [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   257
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   258
lemma leadsTo_weaken_R: "[| F \<in> A leadsTo A'; A'<=B'; st_set(B') |] ==> F \<in> A leadsTo B'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   259
by (blast dest: leadsToD2 intro: subset_imp_leadsTo leadsTo_Trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   260
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   261
lemma leadsTo_weaken_L: "[| F \<in> A leadsTo A'; B<=A |] ==> F \<in> B leadsTo A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   262
apply (frule leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   263
apply (blast intro: leadsTo_Trans subset_imp_leadsTo st_set_subset)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   264
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   265
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   266
lemma leadsTo_weaken: "[| F \<in> A leadsTo A'; B<=A; A'<=B'; st_set(B')|]==> F \<in> B leadsTo B'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   267
apply (frule leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   268
apply (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans leadsTo_refl)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   269
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   270
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   271
(* This rule has a nicer conclusion *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   272
lemma transient_imp_leadsTo2: "[| F \<in> transient(A); state-A<=B; st_set(B)|] ==> F \<in> A leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   273
apply (frule transientD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   274
apply (rule leadsTo_weaken_R)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   275
apply (auto simp add: transient_imp_leadsTo)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   276
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   277
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   278
(*Distributes over binary unions*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   279
lemma leadsTo_Un_distrib: "F:(A Un B) leadsTo C  <->  (F \<in> A leadsTo C & F \<in> B leadsTo C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   280
by (blast intro: leadsTo_Un leadsTo_weaken_L)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   281
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   282
lemma leadsTo_UN_distrib: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   283
"(F:(\<Union>i \<in> I. A(i)) leadsTo B)<-> ((\<forall>i \<in> I. F \<in> A(i) leadsTo B) & F \<in> program & st_set(B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   284
apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   285
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   286
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   287
lemma leadsTo_Union_distrib: "(F \<in> Union(S) leadsTo B) <->  (\<forall>A \<in> S. F \<in> A leadsTo B) & F \<in> program & st_set(B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   288
by (blast dest: leadsToD2 intro: leadsTo_Union leadsTo_weaken_L)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   289
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   290
text{*Set difference: maybe combine with @{text leadsTo_weaken_L}??*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   291
lemma leadsTo_Diff:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   292
     "[| F: (A-B) leadsTo C; F \<in> B leadsTo C; st_set(C) |]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   293
      ==> F \<in> A leadsTo C"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   294
by (blast intro: leadsTo_Un leadsTo_weaken dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   295
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   296
lemma leadsTo_UN_UN:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   297
    "[|!!i. i \<in> I ==> F \<in> A(i) leadsTo A'(i); F \<in> program |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   298
     ==> F: (\<Union>i \<in> I. A(i)) leadsTo (\<Union>i \<in> I. A'(i))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   299
apply (rule leadsTo_Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   300
apply (auto intro: leadsTo_weaken_R dest: leadsToD2) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   301
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   302
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   303
(*Binary union version*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   304
lemma leadsTo_Un_Un: "[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |] ==> F \<in> (A Un B) leadsTo (A' Un B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   305
apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   306
prefer 2 apply (blast dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   307
apply (blast intro: leadsTo_Un leadsTo_weaken_R)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   308
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   309
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   310
(** The cancellation law **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   311
lemma leadsTo_cancel2: "[|F \<in> A leadsTo (A' Un B); F \<in> B leadsTo B'|] ==> F \<in> A leadsTo (A' Un B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   312
apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') &F \<in> program")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   313
prefer 2 apply (blast dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   314
apply (blast intro: leadsTo_Trans leadsTo_Un_Un leadsTo_refl)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   315
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   316
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   317
lemma leadsTo_cancel_Diff2: "[|F \<in> A leadsTo (A' Un B); F \<in> (B-A') leadsTo B'|]==> F \<in> A leadsTo (A' Un B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   318
apply (rule leadsTo_cancel2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   319
prefer 2 apply assumption
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   320
apply (blast dest: leadsToD2 intro: leadsTo_weaken_R)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   321
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   322
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   323
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   324
lemma leadsTo_cancel1: "[| F \<in> A leadsTo (B Un A'); F \<in> B leadsTo B' |] ==> F \<in> A leadsTo (B' Un A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   325
apply (simp add: Un_commute)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   326
apply (blast intro!: leadsTo_cancel2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   327
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   328
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   329
lemma leadsTo_cancel_Diff1:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   330
     "[|F \<in> A leadsTo (B Un A'); F: (B-A') leadsTo B'|]==> F \<in> A leadsTo (B' Un A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   331
apply (rule leadsTo_cancel1)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   332
prefer 2 apply assumption
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   333
apply (blast intro: leadsTo_weaken_R dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   334
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   335
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   336
(*The INDUCTION rule as we should have liked to state it*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   337
lemma leadsTo_induct:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   338
  assumes major: "F \<in> za leadsTo zb"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   339
      and basis: "!!A B. [|F \<in> A ensures B; st_set(A); st_set(B)|] ==> P(A,B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   340
      and trans: "!!A B C. [| F \<in> A leadsTo B; P(A, B);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   341
                              F \<in> B leadsTo C; P(B, C) |] ==> P(A,C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   342
      and union: "!!B S. [| \<forall>A \<in> S. F \<in> A leadsTo B; \<forall>A \<in> S. P(A,B); 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   343
                           st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(Union(S), B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   344
  shows "P(za, zb)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   345
apply (cut_tac major)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   346
apply (unfold leadsTo_def, clarify) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   347
apply (erule leads.induct) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   348
  apply (blast intro: basis [unfolded st_set_def])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   349
 apply (blast intro: trans [unfolded leadsTo_def]) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   350
apply (force intro: union [unfolded st_set_def leadsTo_def]) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   351
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   352
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   353
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   354
(* Added by Sidi, an induction rule without ensures *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   355
lemma leadsTo_induct2:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   356
  assumes major: "F \<in> za leadsTo zb"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   357
      and basis1: "!!A B. [| A<=B; st_set(B) |] ==> P(A, B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   358
      and basis2: "!!A B. [| F \<in> A co A Un B; F \<in> transient(A); st_set(B) |] 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   359
                          ==> P(A, B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   360
      and trans: "!!A B C. [| F \<in> A leadsTo B; P(A, B);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   361
                              F \<in> B leadsTo C; P(B, C) |] ==> P(A,C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   362
      and union: "!!B S. [| \<forall>A \<in> S. F \<in> A leadsTo B; \<forall>A \<in> S. P(A,B); 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   363
                           st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(Union(S), B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   364
  shows "P(za, zb)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   365
apply (cut_tac major)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   366
apply (erule leadsTo_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   367
apply (auto intro: trans union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   368
apply (simp add: ensures_def, clarify)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   369
apply (frule constrainsD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   370
apply (drule_tac B' = " (A-B) Un B" in constrains_weaken_R)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   371
apply blast
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   372
apply (frule ensuresI2 [THEN leadsTo_Basis])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   373
apply (drule_tac [4] basis2, simp_all)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   374
apply (frule_tac A1 = A and B = B in Int_lower2 [THEN basis1])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   375
apply (subgoal_tac "A=Union ({A - B, A Int B}) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   376
prefer 2 apply blast
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   377
apply (erule ssubst)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   378
apply (rule union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   379
apply (auto intro: subset_imp_leadsTo)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   380
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   381
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   382
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   383
(** Variant induction rule: on the preconditions for B **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   384
(*Lemma is the weak version: can't see how to do it in one step*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   385
lemma leadsTo_induct_pre_aux: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   386
  "[| F \<in> za leadsTo zb;   
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   387
      P(zb);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   388
      !!A B. [| F \<in> A ensures B;  P(B); st_set(A); st_set(B) |] ==> P(A);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   389
      !!S. [| \<forall>A \<in> S. P(A); \<forall>A \<in> S. st_set(A) |] ==> P(Union(S))  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   390
   |] ==> P(za)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   391
txt{*by induction on this formula*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   392
apply (subgoal_tac "P (zb) --> P (za) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   393
txt{*now solve first subgoal: this formula is sufficient*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   394
apply (blast intro: leadsTo_refl)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   395
apply (erule leadsTo_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   396
apply (blast+)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   397
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   398
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   399
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   400
lemma leadsTo_induct_pre: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   401
  "[| F \<in> za leadsTo zb;   
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   402
      P(zb);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   403
      !!A B. [| F \<in> A ensures B;  F \<in> B leadsTo zb;  P(B); st_set(A) |] ==> P(A);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   404
      !!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P(A) & st_set(A) ==> P(Union(S))  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   405
   |] ==> P(za)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   406
apply (subgoal_tac " (F \<in> za leadsTo zb) & P (za) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   407
apply (erule conjunct2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   408
apply (frule leadsToD2) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   409
apply (erule leadsTo_induct_pre_aux)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   410
prefer 3 apply (blast dest: leadsToD2 intro: leadsTo_Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   411
prefer 2 apply (blast intro: leadsTo_Trans leadsTo_Basis)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   412
apply (blast intro: leadsTo_refl)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   413
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   414
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   415
(** The impossibility law **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   416
lemma leadsTo_empty: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   417
   "F \<in> A leadsTo 0 ==> A=0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   418
apply (erule leadsTo_induct_pre)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   419
apply (auto simp add: ensures_def constrains_def transient_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   420
apply (drule bspec, assumption)+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   421
apply blast
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   422
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   423
declare leadsTo_empty [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   424
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   425
subsection{*PSP: Progress-Safety-Progress*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   426
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   427
text{*Special case of PSP: Misra's "stable conjunction"*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   428
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   429
lemma psp_stable: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   430
   "[| F \<in> A leadsTo A'; F \<in> stable(B) |] ==> F:(A Int B) leadsTo (A' Int B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   431
apply (unfold stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   432
apply (frule leadsToD2) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   433
apply (erule leadsTo_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   434
prefer 3 apply (blast intro: leadsTo_Union_Int)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   435
prefer 2 apply (blast intro: leadsTo_Trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   436
apply (rule leadsTo_Basis)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   437
apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] Int_Un_distrib2 [symmetric])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   438
apply (auto intro: transient_strengthen constrains_Int)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   439
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   440
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   441
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   442
lemma psp_stable2: "[|F \<in> A leadsTo A'; F \<in> stable(B) |]==>F: (B Int A) leadsTo (B Int A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   443
apply (simp (no_asm_simp) add: psp_stable Int_ac)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   444
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   445
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   446
lemma psp_ensures: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   447
"[| F \<in> A ensures A'; F \<in> B co B' |]==> F: (A Int B') ensures ((A' Int B) Un (B' - B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   448
apply (unfold ensures_def constrains_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   449
(*speeds up the proof*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   450
apply clarify
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   451
apply (blast intro: transient_strengthen)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   452
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   453
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   454
lemma psp: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   455
"[|F \<in> A leadsTo A'; F \<in> B co B'; st_set(B')|]==> F:(A Int B') leadsTo ((A' Int B) Un (B' - B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   456
apply (subgoal_tac "F \<in> program & st_set (A) & st_set (A') & st_set (B) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   457
prefer 2 apply (blast dest!: constrainsD2 leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   458
apply (erule leadsTo_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   459
prefer 3 apply (blast intro: leadsTo_Union_Int)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   460
 txt{*Basis case*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   461
 apply (blast intro: psp_ensures leadsTo_Basis)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   462
txt{*Transitivity case has a delicate argument involving "cancellation"*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   463
apply (rule leadsTo_Un_duplicate2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   464
apply (erule leadsTo_cancel_Diff1)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   465
apply (simp add: Int_Diff Diff_triv)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   466
apply (blast intro: leadsTo_weaken_L dest: constrains_imp_subset)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   467
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   468
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   469
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   470
lemma psp2: "[| F \<in> A leadsTo A'; F \<in> B co B'; st_set(B') |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   471
    ==> F \<in> (B' Int A) leadsTo ((B Int A') Un (B' - B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   472
by (simp (no_asm_simp) add: psp Int_ac)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   473
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   474
lemma psp_unless: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   475
   "[| F \<in> A leadsTo A';  F \<in> B unless B'; st_set(B); st_set(B') |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   476
    ==> F \<in> (A Int B) leadsTo ((A' Int B) Un B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   477
apply (unfold unless_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   478
apply (subgoal_tac "st_set (A) &st_set (A') ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   479
prefer 2 apply (blast dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   480
apply (drule psp, assumption, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   481
apply (blast intro: leadsTo_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   482
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   483
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   484
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   485
subsection{*Proving the induction rules*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   486
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   487
(** The most general rule \<in> r is any wf relation; f is any variant function **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   488
lemma leadsTo_wf_induct_aux: "[| wf(r);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   489
         m \<in> I;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   490
         field(r)<=I;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   491
         F \<in> program; st_set(B); 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   492
         \<forall>m \<in> I. F \<in> (A Int f-``{m}) leadsTo                      
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   493
                    ((A Int f-``(converse(r)``{m})) Un B) |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   494
      ==> F \<in> (A Int f-``{m}) leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   495
apply (erule_tac a = m in wf_induct2, simp_all)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   496
apply (subgoal_tac "F \<in> (A Int (f-`` (converse (r) ``{x}))) leadsTo B")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   497
 apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   498
apply (subst vimage_eq_UN)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   499
apply (simp del: UN_simps add: Int_UN_distrib)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   500
apply (auto intro: leadsTo_UN simp del: UN_simps simp add: Int_UN_distrib)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   501
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   502
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   503
(** Meta or object quantifier ? **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   504
lemma leadsTo_wf_induct: "[| wf(r);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   505
         field(r)<=I;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   506
         A<=f-``I;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   507
         F \<in> program; st_set(A); st_set(B);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   508
         \<forall>m \<in> I. F \<in> (A Int f-``{m}) leadsTo                      
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   509
                    ((A Int f-``(converse(r)``{m})) Un B) |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   510
      ==> F \<in> A leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   511
apply (rule_tac b = A in subst)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   512
 defer 1
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   513
 apply (rule_tac I = I in leadsTo_UN)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   514
 apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   515
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   516
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   517
lemma nat_measure_field: "field(measure(nat, %x. x)) = nat"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   518
apply (unfold field_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   519
apply (simp add: measure_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   520
apply (rule equalityI, force, clarify)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   521
apply (erule_tac V = "x\<notin>range (?y) " in thin_rl)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   522
apply (erule nat_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   523
apply (rule_tac [2] b = "succ (succ (xa))" in domainI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   524
apply (rule_tac b = "succ (0) " in domainI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   525
apply simp_all
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   526
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   527
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   528
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   529
lemma Image_inverse_lessThan: "k<A ==> measure(A, %x. x) -`` {k} = k"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   530
apply (rule equalityI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   531
apply (auto simp add: measure_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   532
apply (blast intro: ltD)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   533
apply (rule vimageI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   534
prefer 2 apply blast
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   535
apply (simp add: lt_Ord lt_Ord2 Ord_mem_iff_lt)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   536
apply (blast intro: lt_trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   537
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   538
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   539
(*Alternative proof is via the lemma F \<in> (A Int f-`(lessThan m)) leadsTo B*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   540
lemma lessThan_induct: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   541
 "[| A<=f-``nat;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   542
     F \<in> program; st_set(A); st_set(B);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   543
     \<forall>m \<in> nat. F:(A Int f-``{m}) leadsTo ((A Int f -`` m) Un B) |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   544
      ==> F \<in> A leadsTo B"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   545
apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct]) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   546
apply (simp_all add: nat_measure_field)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   547
apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   548
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   549
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   550
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   551
(*** wlt ****)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   552
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   553
(*Misra's property W3*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   554
lemma wlt_type: "wlt(F,B) <=state"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   555
by (unfold wlt_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   556
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   557
lemma wlt_st_set: "st_set(wlt(F, B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   558
apply (unfold st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   559
apply (rule wlt_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   560
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   561
declare wlt_st_set [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   562
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   563
lemma wlt_leadsTo_iff: "F \<in> wlt(F, B) leadsTo B <-> (F \<in> program & st_set(B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   564
apply (unfold wlt_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   565
apply (blast dest: leadsToD2 intro!: leadsTo_Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   566
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   567
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   568
(* [| F \<in> program;  st_set(B) |] ==> F \<in> wlt(F, B) leadsTo B  *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   569
lemmas wlt_leadsTo = conjI [THEN wlt_leadsTo_iff [THEN iffD2], standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   570
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   571
lemma leadsTo_subset: "F \<in> A leadsTo B ==> A <= wlt(F, B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   572
apply (unfold wlt_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   573
apply (frule leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   574
apply (auto simp add: st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   575
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   576
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   577
(*Misra's property W2*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   578
lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B <-> (A <= wlt(F,B) & F \<in> program & st_set(B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   579
apply auto
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   580
apply (blast dest: leadsToD2 leadsTo_subset intro: leadsTo_weaken_L wlt_leadsTo)+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   581
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   582
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   583
(*Misra's property W4*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   584
lemma wlt_increasing: "[| F \<in> program; st_set(B) |] ==> B <= wlt(F,B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   585
apply (rule leadsTo_subset)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   586
apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [THEN iff_sym] subset_imp_leadsTo)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   587
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   588
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   589
(*Used in the Trans case below*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   590
lemma leadsTo_123_aux: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   591
   "[| B <= A2;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   592
       F \<in> (A1 - B) co (A1 Un B);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   593
       F \<in> (A2 - C) co (A2 Un C) |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   594
    ==> F \<in> (A1 Un A2 - C) co (A1 Un A2 Un C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   595
apply (unfold constrains_def st_set_def, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   596
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   597
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   598
(*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   599
(* slightly different from the HOL one \<in> B here is bounded *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   600
lemma leadsTo_123: "F \<in> A leadsTo A'  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   601
      ==> \<exists>B \<in> Pow(state). A<=B & F \<in> B leadsTo A' & F \<in> (B-A') co (B Un A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   602
apply (frule leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   603
apply (erule leadsTo_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   604
  txt{*Basis*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   605
  apply (blast dest: ensuresD constrainsD2 st_setD)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   606
 txt{*Trans*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   607
 apply clarify
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   608
 apply (rule_tac x = "Ba Un Bb" in bexI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   609
 apply (blast intro: leadsTo_123_aux leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   610
txt{*Union*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   611
apply (clarify dest!: ball_conj_distrib [THEN iffD1])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   612
apply (subgoal_tac "\<exists>y. y \<in> Pi (S, %A. {Ba \<in> Pow (state) . A<=Ba & F \<in> Ba leadsTo B & F \<in> Ba - B co Ba Un B}) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   613
defer 1
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   614
apply (rule AC_ball_Pi, safe)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   615
apply (rotate_tac 1)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   616
apply (drule_tac x = x in bspec, blast, blast) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   617
apply (rule_tac x = "\<Union>A \<in> S. y`A" in bexI, safe)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   618
apply (rule_tac [3] I1 = S in constrains_UN [THEN constrains_weaken])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   619
apply (rule_tac [2] leadsTo_Union)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   620
prefer 5 apply (blast dest!: apply_type, simp_all)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   621
apply (force dest!: apply_type)+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   622
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   623
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   624
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   625
(*Misra's property W5*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   626
lemma wlt_constrains_wlt: "[| F \<in> program; st_set(B) |] ==>F \<in> (wlt(F, B) - B) co (wlt(F,B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   627
apply (cut_tac F = F in wlt_leadsTo [THEN leadsTo_123], assumption, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   628
apply clarify
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   629
apply (subgoal_tac "Ba = wlt (F,B) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   630
prefer 2 apply (blast dest: leadsTo_eq_subset_wlt [THEN iffD1], clarify)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   631
apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   632
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   633
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   635
subsection{*Completion: Binary and General Finite versions*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   636
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   637
lemma completion_aux: "[| W = wlt(F, (B' Un C));      
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   638
       F \<in> A leadsTo (A' Un C);  F \<in> A' co (A' Un C);    
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   639
       F \<in> B leadsTo (B' Un C);  F \<in> B' co (B' Un C) |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   640
    ==> F \<in> (A Int B) leadsTo ((A' Int B') Un C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   641
apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \<in> program")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   642
 prefer 2 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   643
 apply simp 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   644
 apply (blast dest!: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   645
apply (subgoal_tac "F \<in> (W-C) co (W Un B' Un C) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   646
 prefer 2
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   647
 apply (blast intro!: constrains_weaken [OF constrains_Un [OF _ wlt_constrains_wlt]])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   648
apply (subgoal_tac "F \<in> (W-C) co W")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   649
 prefer 2
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   650
 apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]] Un_assoc)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   651
apply (subgoal_tac "F \<in> (A Int W - C) leadsTo (A' Int W Un C) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   652
 prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   653
(** step 13 **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   654
apply (subgoal_tac "F \<in> (A' Int W Un C) leadsTo (A' Int B' Un C) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   655
apply (drule leadsTo_Diff)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   656
apply (blast intro: subset_imp_leadsTo dest: leadsToD2 constrainsD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   657
apply (force simp add: st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   658
apply (subgoal_tac "A Int B <= A Int W")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   659
prefer 2 apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   660
apply (blast intro: leadsTo_Trans subset_imp_leadsTo)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   661
txt{*last subgoal*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   662
apply (rule_tac leadsTo_Un_duplicate2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   663
apply (rule_tac leadsTo_Un_Un)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   664
 prefer 2 apply (blast intro: leadsTo_refl)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   665
apply (rule_tac A'1 = "B' Un C" in wlt_leadsTo[THEN psp2, THEN leadsTo_weaken])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   666
apply blast+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   667
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   668
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   669
lemmas completion = refl [THEN completion_aux, standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   670
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   671
lemma finite_completion_aux:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   672
     "[| I \<in> Fin(X); F \<in> program; st_set(C) |] ==>  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   673
       (\<forall>i \<in> I. F \<in> (A(i)) leadsTo (A'(i) Un C)) -->   
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   674
                     (\<forall>i \<in> I. F \<in> (A'(i)) co (A'(i) Un C)) -->  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   675
                   F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   676
apply (erule Fin_induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   677
apply (auto simp add: Inter_0)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   678
apply (rule completion)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   679
apply (auto simp del: INT_simps simp add: INT_extend_simps)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   680
apply (blast intro: constrains_INT)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   681
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   682
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   683
lemma finite_completion: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   684
     "[| I \<in> Fin(X);   
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   685
         !!i. i \<in> I ==> F \<in> A(i) leadsTo (A'(i) Un C);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   686
         !!i. i \<in> I ==> F \<in> A'(i) co (A'(i) Un C); F \<in> program; st_set(C)|]    
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   687
      ==> F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   688
by (blast intro: finite_completion_aux [THEN mp, THEN mp])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   689
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   690
lemma stable_completion: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   691
     "[| F \<in> A leadsTo A';  F \<in> stable(A');    
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   692
         F \<in> B leadsTo B';  F \<in> stable(B') |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   693
    ==> F \<in> (A Int B) leadsTo (A' Int B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   694
apply (unfold stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   695
apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   696
apply (blast dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   697
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   698
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   699
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   700
lemma finite_stable_completion: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   701
     "[| I \<in> Fin(X);  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   702
         (!!i. i \<in> I ==> F \<in> A(i) leadsTo A'(i));  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   703
         (!!i. i \<in> I ==> F \<in> stable(A'(i)));  F \<in> program |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   704
      ==> F \<in> (\<Inter>i \<in> I. A(i)) leadsTo (\<Inter>i \<in> I. A'(i))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   705
apply (unfold stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   706
apply (subgoal_tac "st_set (\<Inter>i \<in> I. A' (i))")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   707
prefer 2 apply (blast dest: leadsToD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   708
apply (rule_tac C1 = 0 in finite_completion [THEN leadsTo_weaken_R], auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   709
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   710
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   711
ML
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   712
{*
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   713
val Int_Union_Union = thm "Int_Union_Union";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   714
val Int_Union_Union2 = thm "Int_Union_Union2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   715
val transient_type = thm "transient_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   716
val transientD2 = thm "transientD2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   717
val stable_transient_empty = thm "stable_transient_empty";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   718
val transient_strengthen = thm "transient_strengthen";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   719
val transientI = thm "transientI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   720
val transientE = thm "transientE";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   721
val transient_state = thm "transient_state";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   722
val transient_state2 = thm "transient_state2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   723
val transient_empty = thm "transient_empty";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   724
val ensures_type = thm "ensures_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   725
val ensuresI = thm "ensuresI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   726
val ensuresI2 = thm "ensuresI2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   727
val ensuresD = thm "ensuresD";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   728
val ensures_weaken_R = thm "ensures_weaken_R";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   729
val stable_ensures_Int = thm "stable_ensures_Int";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   730
val stable_transient_ensures = thm "stable_transient_ensures";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   731
val ensures_eq = thm "ensures_eq";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   732
val subset_imp_ensures = thm "subset_imp_ensures";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   733
val leads_left = thm "leads_left";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   734
val leads_right = thm "leads_right";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   735
val leadsTo_type = thm "leadsTo_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   736
val leadsToD2 = thm "leadsToD2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   737
val leadsTo_Basis = thm "leadsTo_Basis";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   738
val subset_imp_leadsTo = thm "subset_imp_leadsTo";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   739
val leadsTo_Trans = thm "leadsTo_Trans";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   740
val transient_imp_leadsTo = thm "transient_imp_leadsTo";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   741
val leadsTo_Un_duplicate = thm "leadsTo_Un_duplicate";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   742
val leadsTo_Un_duplicate2 = thm "leadsTo_Un_duplicate2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   743
val leadsTo_Union = thm "leadsTo_Union";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   744
val leadsTo_Union_Int = thm "leadsTo_Union_Int";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   745
val leadsTo_UN = thm "leadsTo_UN";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   746
val leadsTo_Un = thm "leadsTo_Un";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   747
val single_leadsTo_I = thm "single_leadsTo_I";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   748
val leadsTo_refl = thm "leadsTo_refl";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   749
val leadsTo_refl_iff = thm "leadsTo_refl_iff";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   750
val empty_leadsTo = thm "empty_leadsTo";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   751
val leadsTo_state = thm "leadsTo_state";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   752
val leadsTo_weaken_R = thm "leadsTo_weaken_R";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   753
val leadsTo_weaken_L = thm "leadsTo_weaken_L";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   754
val leadsTo_weaken = thm "leadsTo_weaken";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   755
val transient_imp_leadsTo2 = thm "transient_imp_leadsTo2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   756
val leadsTo_Un_distrib = thm "leadsTo_Un_distrib";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   757
val leadsTo_UN_distrib = thm "leadsTo_UN_distrib";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   758
val leadsTo_Union_distrib = thm "leadsTo_Union_distrib";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   759
val leadsTo_Diff = thm "leadsTo_Diff";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   760
val leadsTo_UN_UN = thm "leadsTo_UN_UN";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   761
val leadsTo_Un_Un = thm "leadsTo_Un_Un";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   762
val leadsTo_cancel2 = thm "leadsTo_cancel2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   763
val leadsTo_cancel_Diff2 = thm "leadsTo_cancel_Diff2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   764
val leadsTo_cancel1 = thm "leadsTo_cancel1";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   765
val leadsTo_cancel_Diff1 = thm "leadsTo_cancel_Diff1";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   766
val leadsTo_induct = thm "leadsTo_induct";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   767
val leadsTo_induct2 = thm "leadsTo_induct2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   768
val leadsTo_induct_pre_aux = thm "leadsTo_induct_pre_aux";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   769
val leadsTo_induct_pre = thm "leadsTo_induct_pre";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   770
val leadsTo_empty = thm "leadsTo_empty";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   771
val psp_stable = thm "psp_stable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   772
val psp_stable2 = thm "psp_stable2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   773
val psp_ensures = thm "psp_ensures";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   774
val psp = thm "psp";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   775
val psp2 = thm "psp2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   776
val psp_unless = thm "psp_unless";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   777
val leadsTo_wf_induct_aux = thm "leadsTo_wf_induct_aux";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   778
val leadsTo_wf_induct = thm "leadsTo_wf_induct";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   779
val nat_measure_field = thm "nat_measure_field";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   780
val Image_inverse_lessThan = thm "Image_inverse_lessThan";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   781
val lessThan_induct = thm "lessThan_induct";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   782
val wlt_type = thm "wlt_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   783
val wlt_st_set = thm "wlt_st_set";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   784
val wlt_leadsTo_iff = thm "wlt_leadsTo_iff";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   785
val wlt_leadsTo = thm "wlt_leadsTo";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   786
val leadsTo_subset = thm "leadsTo_subset";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   787
val leadsTo_eq_subset_wlt = thm "leadsTo_eq_subset_wlt";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   788
val wlt_increasing = thm "wlt_increasing";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   789
val leadsTo_123_aux = thm "leadsTo_123_aux";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   790
val leadsTo_123 = thm "leadsTo_123";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   791
val wlt_constrains_wlt = thm "wlt_constrains_wlt";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   792
val completion_aux = thm "completion_aux";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   793
val completion = thm "completion";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   794
val finite_completion_aux = thm "finite_completion_aux";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   795
val finite_completion = thm "finite_completion";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   796
val stable_completion = thm "stable_completion";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   797
val finite_stable_completion = thm "finite_stable_completion";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14077
diff changeset
   798
*}
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   799
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   800
end