src/HOL/UNITY/WFair.thy
author wenzelm
Tue, 16 Jan 2018 09:30:00 +0100
changeset 67443 3abf6a722518
parent 63146 f1ecba0272f9
child 67613 ce654b0e6d69
permissions -rw-r--r--
standardized towards new-style formal comments: isabelle update_comments;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 35422
diff changeset
     1
(*  Title:      HOL/UNITY/WFair.thy
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
     5
Conditional Fairness versions of transient, ensures, leadsTo.
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
From Misra, "A Logic for Concurrent Programming", 1994
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    10
section\<open>Progress\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
    11
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15045
diff changeset
    12
theory WFair imports UNITY begin
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    14
text\<open>The original version of this theory was based on weak fairness.  (Thus,
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    15
the entire UNITY development embodied this assumption, until February 2003.)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    16
Weak fairness states that if a command is enabled continuously, then it is
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    17
eventually executed.  Ernie Cohen suggested that I instead adopt unconditional
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    18
fairness: every command is executed infinitely often.  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    19
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    20
In fact, Misra's paper on "Progress" seems to be ambiguous about the correct
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    21
interpretation, and says that the two forms of fairness are equivalent.  They
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    22
differ only on their treatment of partial transitions, which under
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    23
unconditional fairness behave magically.  That is because if there are partial
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    24
transitions then there may be no fair executions, making all leads-to
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    25
properties hold vacuously.
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    26
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    27
Unconditional fairness has some great advantages.  By distinguishing partial
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    28
transitions from total ones that are the identity on part of their domain, it
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    29
is more expressive.  Also, by simplifying the definition of the transient
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    30
property, it simplifies many proofs.  A drawback is that some laws only hold
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    31
under the assumption that all transitions are total.  The best-known of these
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    32
is the impossibility law for leads-to.
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    33
\<close>
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    34
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32693
diff changeset
    35
definition
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    36
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
    37
  \<comment> \<open>This definition specifies conditional fairness.  The rest of the theory
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    38
      is generic to all forms of fairness.  To get weak fairness, conjoin
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    39
      the inclusion below with @{term "A \<subseteq> Domain act"}, which specifies 
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    40
      that the action is enabled over all of @{term A}.\<close>
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32693
diff changeset
    41
  transient :: "'a set => 'a program set" where
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    42
    "transient A == {F. \<exists>act\<in>Acts F. act``A \<subseteq> -A}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32693
diff changeset
    44
definition
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32693
diff changeset
    45
  ensures :: "['a set, 'a set] => 'a program set"       (infixl "ensures" 60) where
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    46
    "A ensures B == (A-B co A \<union> B) \<inter> transient (A-B)"
8006
paulson
parents: 7346
diff changeset
    47
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    48
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 19769
diff changeset
    49
inductive_set
6801
9e0037839d63 renamed the underlying relation of leadsTo from "leadsto"
paulson
parents: 6536
diff changeset
    50
  leads :: "'a program => ('a set * 'a set) set"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
    51
    \<comment> \<open>LEADS-TO constant for the inductive definition\<close>
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 19769
diff changeset
    52
  for F :: "'a program"
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 19769
diff changeset
    53
  where
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    55
    Basis:  "F \<in> A ensures B ==> (A,B) \<in> leads F"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 19769
diff changeset
    57
  | Trans:  "[| (A,B) \<in> leads F;  (B,C) \<in> leads F |] ==> (A,C) \<in> leads F"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    58
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 19769
diff changeset
    59
  | Union:  "\<forall>A \<in> S. (A,B) \<in> leads F ==> (Union S, B) \<in> leads F"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    60
5155
21177b8a4d7f added comments
paulson
parents: 4776
diff changeset
    61
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32693
diff changeset
    62
definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
    63
     \<comment> \<open>visible version of the LEADS-TO relation\<close>
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    64
    "A leadsTo B == {F. (A,B) \<in> leads F}"
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5340
diff changeset
    65
  
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32693
diff changeset
    66
definition wlt :: "['a program, 'a set] => 'a set" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
    67
     \<comment> \<open>predicate transformer: the largest set that leads to @{term B}\<close>
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
    68
    "wlt F B == \<Union>{A. F \<in> A leadsTo B}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    69
60773
d09c66a0ea10 more symbols by default, without xsymbols mode;
wenzelm
parents: 58889
diff changeset
    70
notation leadsTo  (infixl "\<longmapsto>" 60)
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    71
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    72
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    73
subsection\<open>transient\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    74
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    75
lemma stable_transient: 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    76
    "[| F \<in> stable A; F \<in> transient A |] ==> \<exists>act\<in>Acts F. A \<subseteq> - (Domain act)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    77
apply (simp add: stable_def constrains_def transient_def, clarify)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    78
apply (rule rev_bexI, auto)  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    79
done
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    80
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    81
lemma stable_transient_empty: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    82
    "[| F \<in> stable A; F \<in> transient A; all_total F |] ==> A = {}"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    83
apply (drule stable_transient, assumption)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    84
apply (simp add: all_total_def)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    85
done
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    86
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    87
lemma transient_strengthen: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    88
    "[| F \<in> transient A; B \<subseteq> A |] ==> F \<in> transient B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    89
apply (unfold transient_def, clarify)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    90
apply (blast intro!: rev_bexI)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    91
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    92
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    93
lemma transientI: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    94
    "[| act: Acts F;  act``A \<subseteq> -A |] ==> F \<in> transient A"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    95
by (unfold transient_def, blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    96
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
    97
lemma transientE: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    98
    "[| F \<in> transient A;   
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
    99
        !!act. [| act: Acts F;  act``A \<subseteq> -A |] ==> P |]  
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   100
     ==> P"
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   101
by (unfold transient_def, blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   102
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   103
lemma transient_empty [simp]: "transient {} = UNIV"
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   104
by (unfold transient_def, auto)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   105
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   106
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   107
text\<open>This equation recovers the notion of weak fairness.  A totalized
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   108
      program satisfies a transient assertion just if the original program
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   109
      contains a suitable action that is also enabled.\<close>
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   110
lemma totalize_transient_iff:
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   111
   "(totalize F \<in> transient A) = (\<exists>act\<in>Acts F. A \<subseteq> Domain act & act``A \<subseteq> -A)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   112
apply (simp add: totalize_def totalize_act_def transient_def 
32693
6c6b1ba5e71e tuned proofs
haftmann
parents: 23767
diff changeset
   113
                 Un_Image, safe)
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   114
apply (blast intro!: rev_bexI)+
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   115
done
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   116
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   117
lemma totalize_transientI: 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   118
    "[| act: Acts F;  A \<subseteq> Domain act;  act``A \<subseteq> -A |] 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   119
     ==> totalize F \<in> transient A"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   120
by (simp add: totalize_transient_iff, blast)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   121
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   122
subsection\<open>ensures\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   123
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   124
lemma ensuresI: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   125
    "[| F \<in> (A-B) co (A \<union> B); F \<in> transient (A-B) |] ==> F \<in> A ensures B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   126
by (unfold ensures_def, blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   127
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   128
lemma ensuresD: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   129
    "F \<in> A ensures B ==> F \<in> (A-B) co (A \<union> B) & F \<in> transient (A-B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   130
by (unfold ensures_def, blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   131
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   132
lemma ensures_weaken_R: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   133
    "[| F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   134
apply (unfold ensures_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   135
apply (blast intro: constrains_weaken transient_strengthen)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   136
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   137
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   138
text\<open>The L-version (precondition strengthening) fails, but we have this\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   139
lemma stable_ensures_Int: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   140
    "[| F \<in> stable C;  F \<in> A ensures B |]    
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   141
    ==> F \<in> (C \<inter> A) ensures (C \<inter> B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   142
apply (unfold ensures_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   143
apply (auto simp add: ensures_def Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   144
prefer 2 apply (blast intro: transient_strengthen)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   145
apply (blast intro: stable_constrains_Int constrains_weaken)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   146
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   147
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   148
lemma stable_transient_ensures:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   149
     "[| F \<in> stable A;  F \<in> transient C;  A \<subseteq> B \<union> C |] ==> F \<in> A ensures B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   150
apply (simp add: ensures_def stable_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   151
apply (blast intro: constrains_weaken transient_strengthen)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   152
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   153
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   154
lemma ensures_eq: "(A ensures B) = (A unless B) \<inter> transient (A-B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   155
by (simp (no_asm) add: ensures_def unless_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   156
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   157
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   158
subsection\<open>leadsTo\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   159
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   160
lemma leadsTo_Basis [intro]: "F \<in> A ensures B ==> F \<in> A leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   161
apply (unfold leadsTo_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   162
apply (blast intro: leads.Basis)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   163
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   164
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   165
lemma leadsTo_Trans: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   166
     "[| F \<in> A leadsTo B;  F \<in> B leadsTo C |] ==> F \<in> A leadsTo C"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   167
apply (unfold leadsTo_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   168
apply (blast intro: leads.Trans)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   169
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   170
14112
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   171
lemma leadsTo_Basis':
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   172
     "[| F \<in> A co A \<union> B; F \<in> transient A |] ==> F \<in> A leadsTo B"
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   173
apply (drule_tac B = "A-B" in constrains_weaken_L)
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   174
apply (drule_tac [2] B = "A-B" in transient_strengthen)
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   175
apply (rule_tac [3] ensuresI [THEN leadsTo_Basis])
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   176
apply (blast+)
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   177
done
95d51043d2a3 tidying
paulson
parents: 13812
diff changeset
   178
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   179
lemma transient_imp_leadsTo: "F \<in> transient A ==> F \<in> A leadsTo (-A)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   180
by (simp (no_asm_simp) add: leadsTo_Basis ensuresI Compl_partition)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   181
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   182
text\<open>Useful with cancellation, disjunction\<close>
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   183
lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' \<union> A') ==> F \<in> A leadsTo A'"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   184
by (simp add: Un_ac)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   185
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   186
lemma leadsTo_Un_duplicate2:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   187
     "F \<in> A leadsTo (A' \<union> C \<union> C) ==> F \<in> A leadsTo (A' \<union> C)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   188
by (simp add: Un_ac)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   189
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   190
text\<open>The Union introduction rule as we should have liked to state it\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   191
lemma leadsTo_Union: 
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
   192
    "(!!A. A \<in> S ==> F \<in> A leadsTo B) ==> F \<in> (\<Union>S) leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   193
apply (unfold leadsTo_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   194
apply (blast intro: leads.Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   195
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   196
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   197
lemma leadsTo_Union_Int: 
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
   198
 "(!!A. A \<in> S ==> F \<in> (A \<inter> C) leadsTo B) ==> F \<in> (\<Union>S \<inter> C) leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   199
apply (unfold leadsTo_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   200
apply (simp only: Int_Union_Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   201
apply (blast intro: leads.Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   202
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   203
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   204
lemma leadsTo_UN: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   205
    "(!!i. i \<in> I ==> F \<in> (A i) leadsTo B) ==> F \<in> (\<Union>i \<in> I. A i) leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   206
apply (blast intro: leadsTo_Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   207
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   208
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   209
text\<open>Binary union introduction rule\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   210
lemma leadsTo_Un:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   211
     "[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A \<union> B) leadsTo C"
44106
0e018cbcc0de tuned proofs
haftmann
parents: 37936
diff changeset
   212
  using leadsTo_Union [of "{A, B}" F C] by auto
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   213
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   214
lemma single_leadsTo_I: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   215
     "(!!x. x \<in> A ==> F \<in> {x} leadsTo B) ==> F \<in> A leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   216
by (subst UN_singleton [symmetric], rule leadsTo_UN, blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   217
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   218
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   219
text\<open>The INDUCTION rule as we should have liked to state it\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   220
lemma leadsTo_induct: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   221
  "[| F \<in> za leadsTo zb;   
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   222
      !!A B. F \<in> A ensures B ==> P A B;  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   223
      !!A B C. [| F \<in> A leadsTo B; P A B; F \<in> B leadsTo C; P B C |]  
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   224
               ==> P A C;  
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
   225
      !!B S. \<forall>A \<in> S. F \<in> A leadsTo B & P A B ==> P (\<Union>S) B  
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   226
   |] ==> P za zb"
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   227
apply (unfold leadsTo_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   228
apply (drule CollectD, erule leads.induct)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   229
apply (blast+)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   230
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   231
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   232
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   233
lemma subset_imp_ensures: "A \<subseteq> B ==> F \<in> A ensures B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   234
by (unfold ensures_def constrains_def transient_def, blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   235
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 45477
diff changeset
   236
lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis]
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   237
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 45477
diff changeset
   238
lemmas leadsTo_refl = subset_refl [THEN subset_imp_leadsTo]
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   239
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 45477
diff changeset
   240
lemmas empty_leadsTo = empty_subsetI [THEN subset_imp_leadsTo, simp]
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   241
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 45477
diff changeset
   242
lemmas leadsTo_UNIV = subset_UNIV [THEN subset_imp_leadsTo, simp]
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   243
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   244
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   245
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   246
(** Variant induction rule: on the preconditions for B **)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   247
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   248
text\<open>Lemma is the weak version: can't see how to do it in one step\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   249
lemma leadsTo_induct_pre_lemma: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   250
  "[| F \<in> za leadsTo zb;   
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   251
      P zb;  
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   252
      !!A B. [| F \<in> A ensures B;  P B |] ==> P A;  
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
   253
      !!S. \<forall>A \<in> S. P A ==> P (\<Union>S)  
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   254
   |] ==> P za"
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   255
txt\<open>by induction on this formula\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   256
apply (subgoal_tac "P zb --> P za")
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   257
txt\<open>now solve first subgoal: this formula is sufficient\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   258
apply (blast intro: leadsTo_refl)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   259
apply (erule leadsTo_induct)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   260
apply (blast+)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   261
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   262
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   263
lemma leadsTo_induct_pre: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   264
  "[| F \<in> za leadsTo zb;   
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   265
      P zb;  
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   266
      !!A B. [| F \<in> A ensures B;  F \<in> B leadsTo zb;  P B |] ==> P A;  
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
   267
      !!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P A ==> P (\<Union>S)  
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   268
   |] ==> P za"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   269
apply (subgoal_tac "F \<in> za leadsTo zb & P za")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   270
apply (erule conjunct2)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   271
apply (erule leadsTo_induct_pre_lemma)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   272
prefer 3 apply (blast intro: leadsTo_Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   273
prefer 2 apply (blast intro: leadsTo_Trans)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   274
apply (blast intro: leadsTo_refl)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   275
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   276
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   277
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   278
lemma leadsTo_weaken_R: "[| F \<in> A leadsTo A'; A'<=B' |] ==> F \<in> A leadsTo B'"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   279
by (blast intro: subset_imp_leadsTo leadsTo_Trans)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   280
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 44106
diff changeset
   281
lemma leadsTo_weaken_L:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   282
     "[| F \<in> A leadsTo A'; B \<subseteq> A |] ==> F \<in> B leadsTo A'"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   283
by (blast intro: leadsTo_Trans subset_imp_leadsTo)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   284
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   285
text\<open>Distributes over binary unions\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   286
lemma leadsTo_Un_distrib:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   287
     "F \<in> (A \<union> B) leadsTo C  =  (F \<in> A leadsTo C & F \<in> B leadsTo C)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   288
by (blast intro: leadsTo_Un leadsTo_weaken_L)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   289
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   290
lemma leadsTo_UN_distrib:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   291
     "F \<in> (\<Union>i \<in> I. A i) leadsTo B  =  (\<forall>i \<in> I. F \<in> (A i) leadsTo B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   292
by (blast intro: leadsTo_UN leadsTo_weaken_L)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   293
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   294
lemma leadsTo_Union_distrib:
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61824
diff changeset
   295
     "F \<in> (\<Union>S) leadsTo B  =  (\<forall>A \<in> S. F \<in> A leadsTo B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   296
by (blast intro: leadsTo_Union leadsTo_weaken_L)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   297
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   298
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   299
lemma leadsTo_weaken:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   300
     "[| F \<in> A leadsTo A'; B \<subseteq> A; A'<=B' |] ==> F \<in> B leadsTo B'"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   301
by (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   302
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   303
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   304
text\<open>Set difference: maybe combine with \<open>leadsTo_weaken_L\<close>??\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   305
lemma leadsTo_Diff:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   306
     "[| F \<in> (A-B) leadsTo C; F \<in> B leadsTo C |]   ==> F \<in> A leadsTo C"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   307
by (blast intro: leadsTo_Un leadsTo_weaken)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   308
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   309
lemma leadsTo_UN_UN:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   310
   "(!! i. i \<in> I ==> F \<in> (A i) leadsTo (A' i))  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   311
    ==> F \<in> (\<Union>i \<in> I. A i) leadsTo (\<Union>i \<in> I. A' i)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   312
apply (blast intro: leadsTo_Union leadsTo_weaken_R)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   313
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   314
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   315
text\<open>Binary union version\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   316
lemma leadsTo_Un_Un:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   317
     "[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   318
      ==> F \<in> (A \<union> B) leadsTo (A' \<union> B')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   319
by (blast intro: leadsTo_Un leadsTo_weaken_R)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   320
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   321
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   322
(** The cancellation law **)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   323
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   324
lemma leadsTo_cancel2:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   325
     "[| F \<in> A leadsTo (A' \<union> B); F \<in> B leadsTo B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   326
      ==> F \<in> A leadsTo (A' \<union> B')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   327
by (blast intro: leadsTo_Un_Un subset_imp_leadsTo leadsTo_Trans)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   328
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   329
lemma leadsTo_cancel_Diff2:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   330
     "[| F \<in> A leadsTo (A' \<union> B); F \<in> (B-A') leadsTo B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   331
      ==> F \<in> A leadsTo (A' \<union> B')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   332
apply (rule leadsTo_cancel2)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   333
prefer 2 apply assumption
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   334
apply (simp_all (no_asm_simp))
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   335
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   336
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   337
lemma leadsTo_cancel1:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   338
     "[| F \<in> A leadsTo (B \<union> A'); F \<in> B leadsTo B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   339
    ==> F \<in> A leadsTo (B' \<union> A')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   340
apply (simp add: Un_commute)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   341
apply (blast intro!: leadsTo_cancel2)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   342
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   343
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   344
lemma leadsTo_cancel_Diff1:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   345
     "[| F \<in> A leadsTo (B \<union> A'); F \<in> (B-A') leadsTo B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   346
    ==> F \<in> A leadsTo (B' \<union> A')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   347
apply (rule leadsTo_cancel1)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   348
prefer 2 apply assumption
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   349
apply (simp_all (no_asm_simp))
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   350
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   351
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   352
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   353
text\<open>The impossibility law\<close>
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   354
lemma leadsTo_empty: "[|F \<in> A leadsTo {}; all_total F|] ==> A={}"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   355
apply (erule leadsTo_induct_pre)
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   356
apply (simp_all add: ensures_def constrains_def transient_def all_total_def, clarify)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   357
apply (drule bspec, assumption)+
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   358
apply blast
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   359
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   360
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   361
subsection\<open>PSP: Progress-Safety-Progress\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   362
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   363
text\<open>Special case of PSP: Misra's "stable conjunction"\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   364
lemma psp_stable: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   365
   "[| F \<in> A leadsTo A'; F \<in> stable B |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   366
    ==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   367
apply (unfold stable_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   368
apply (erule leadsTo_induct)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   369
prefer 3 apply (blast intro: leadsTo_Union_Int)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   370
prefer 2 apply (blast intro: leadsTo_Trans)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   371
apply (rule leadsTo_Basis)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   372
apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] Int_Un_distrib2 [symmetric])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   373
apply (blast intro: transient_strengthen constrains_Int)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   374
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   375
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   376
lemma psp_stable2: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   377
   "[| F \<in> A leadsTo A'; F \<in> stable B |] ==> F \<in> (B \<inter> A) leadsTo (B \<inter> A')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   378
by (simp add: psp_stable Int_ac)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   379
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   380
lemma psp_ensures: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   381
   "[| F \<in> A ensures A'; F \<in> B co B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   382
    ==> F \<in> (A \<inter> B') ensures ((A' \<inter> B) \<union> (B' - B))"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   383
apply (unfold ensures_def constrains_def, clarify) (*speeds up the proof*)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   384
apply (blast intro: transient_strengthen)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   385
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   386
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   387
lemma psp:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   388
     "[| F \<in> A leadsTo A'; F \<in> B co B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   389
      ==> F \<in> (A \<inter> B') leadsTo ((A' \<inter> B) \<union> (B' - B))"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   390
apply (erule leadsTo_induct)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   391
  prefer 3 apply (blast intro: leadsTo_Union_Int)
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   392
 txt\<open>Basis case\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   393
 apply (blast intro: psp_ensures)
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   394
txt\<open>Transitivity case has a delicate argument involving "cancellation"\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   395
apply (rule leadsTo_Un_duplicate2)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   396
apply (erule leadsTo_cancel_Diff1)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   397
apply (simp add: Int_Diff Diff_triv)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   398
apply (blast intro: leadsTo_weaken_L dest: constrains_imp_subset)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   399
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   400
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   401
lemma psp2:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   402
     "[| F \<in> A leadsTo A'; F \<in> B co B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   403
    ==> F \<in> (B' \<inter> A) leadsTo ((B \<inter> A') \<union> (B' - B))"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   404
by (simp (no_asm_simp) add: psp Int_ac)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   405
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   406
lemma psp_unless: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   407
   "[| F \<in> A leadsTo A';  F \<in> B unless B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   408
    ==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B) \<union> B')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   409
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   410
apply (unfold unless_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   411
apply (drule psp, assumption)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   412
apply (blast intro: leadsTo_weaken)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   413
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   414
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   415
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   416
subsection\<open>Proving the induction rules\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   417
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   418
(** The most general rule: r is any wf relation; f is any variant function **)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   419
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   420
lemma leadsTo_wf_induct_lemma:
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   421
     "[| wf r;      
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   422
         \<forall>m. F \<in> (A \<inter> f-`{m}) leadsTo                      
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   423
                    ((A \<inter> f-`(r^-1 `` {m})) \<union> B) |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   424
      ==> F \<in> (A \<inter> f-`{m}) leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   425
apply (erule_tac a = m in wf_induct)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   426
apply (subgoal_tac "F \<in> (A \<inter> (f -` (r^-1 `` {x}))) leadsTo B")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   427
 apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   428
apply (subst vimage_eq_UN)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   429
apply (simp only: UN_simps [symmetric])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   430
apply (blast intro: leadsTo_UN)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   431
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   432
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   433
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   434
(** Meta or object quantifier ? **)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   435
lemma leadsTo_wf_induct:
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   436
     "[| wf r;      
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   437
         \<forall>m. F \<in> (A \<inter> f-`{m}) leadsTo                      
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   438
                    ((A \<inter> f-`(r^-1 `` {m})) \<union> B) |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   439
      ==> F \<in> A leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   440
apply (rule_tac t = A in subst)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   441
 defer 1
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   442
 apply (rule leadsTo_UN)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   443
 apply (erule leadsTo_wf_induct_lemma)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   444
 apply assumption
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   445
apply fast (*Blast_tac: Function unknown's argument not a parameter*)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   446
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   447
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   448
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   449
lemma bounded_induct:
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   450
     "[| wf r;      
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   451
         \<forall>m \<in> I. F \<in> (A \<inter> f-`{m}) leadsTo                    
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   452
                      ((A \<inter> f-`(r^-1 `` {m})) \<union> B) |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   453
      ==> F \<in> A leadsTo ((A - (f-`I)) \<union> B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   454
apply (erule leadsTo_wf_induct, safe)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   455
apply (case_tac "m \<in> I")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   456
apply (blast intro: leadsTo_weaken)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   457
apply (blast intro: subset_imp_leadsTo)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   458
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   459
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   460
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   461
(*Alternative proof is via the lemma F \<in> (A \<inter> f-`(lessThan m)) leadsTo B*)
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   462
lemma lessThan_induct: 
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14150
diff changeset
   463
     "[| !!m::nat. F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`{..<m}) \<union> B) |]  
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   464
      ==> F \<in> A leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   465
apply (rule wf_less_than [THEN leadsTo_wf_induct])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   466
apply (simp (no_asm_simp))
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   467
apply blast
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   468
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   469
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   470
lemma lessThan_bounded_induct:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   471
     "!!l::nat. [| \<forall>m \<in> greaterThan l.     
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   472
            F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`(lessThan m)) \<union> B) |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   473
      ==> F \<in> A leadsTo ((A \<inter> (f-`(atMost l))) \<union> B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   474
apply (simp only: Diff_eq [symmetric] vimage_Compl Compl_greaterThan [symmetric])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   475
apply (rule wf_less_than [THEN bounded_induct])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   476
apply (simp (no_asm_simp))
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   477
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   478
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   479
lemma greaterThan_bounded_induct:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   480
     "(!!l::nat. \<forall>m \<in> lessThan l.     
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   481
                 F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`(greaterThan m)) \<union> B))
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   482
      ==> F \<in> A leadsTo ((A \<inter> (f-`(atLeast l))) \<union> B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   483
apply (rule_tac f = f and f1 = "%k. l - k" 
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   484
       in wf_less_than [THEN wf_inv_image, THEN leadsTo_wf_induct])
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 16417
diff changeset
   485
apply (simp (no_asm) add:Image_singleton)
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   486
apply clarify
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   487
apply (case_tac "m<l")
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   488
 apply (blast intro: leadsTo_weaken_R diff_less_mono2)
61824
dcbe9f756ae0 not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents: 60773
diff changeset
   489
apply (blast intro: not_le_imp_less subset_imp_leadsTo)
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   490
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   491
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   492
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   493
subsection\<open>wlt\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   494
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   495
text\<open>Misra's property W3\<close>
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   496
lemma wlt_leadsTo: "F \<in> (wlt F B) leadsTo B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   497
apply (unfold wlt_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   498
apply (blast intro!: leadsTo_Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   499
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   500
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   501
lemma leadsTo_subset: "F \<in> A leadsTo B ==> A \<subseteq> wlt F B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   502
apply (unfold wlt_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   503
apply (blast intro!: leadsTo_Union)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   504
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   505
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   506
text\<open>Misra's property W2\<close>
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   507
lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B = (A \<subseteq> wlt F B)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   508
by (blast intro!: leadsTo_subset wlt_leadsTo [THEN leadsTo_weaken_L])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   509
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   510
text\<open>Misra's property W4\<close>
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   511
lemma wlt_increasing: "B \<subseteq> wlt F B"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   512
apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [symmetric] subset_imp_leadsTo)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   513
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   514
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   515
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   516
text\<open>Used in the Trans case below\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   517
lemma lemma1: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   518
   "[| B \<subseteq> A2;   
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   519
       F \<in> (A1 - B) co (A1 \<union> B);  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   520
       F \<in> (A2 - C) co (A2 \<union> C) |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   521
    ==> F \<in> (A1 \<union> A2 - C) co (A1 \<union> A2 \<union> C)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   522
by (unfold constrains_def, clarify,  blast)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   523
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   524
text\<open>Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   525
lemma leadsTo_123:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   526
     "F \<in> A leadsTo A'  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   527
      ==> \<exists>B. A \<subseteq> B & F \<in> B leadsTo A' & F \<in> (B-A') co (B \<union> A')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   528
apply (erule leadsTo_induct)
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   529
  txt\<open>Basis\<close>
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   530
  apply (blast dest: ensuresD)
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   531
 txt\<open>Trans\<close>
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   532
 apply clarify
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   533
 apply (rule_tac x = "Ba \<union> Bb" in exI)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   534
 apply (blast intro: lemma1 leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate)
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   535
txt\<open>Union\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   536
apply (clarify dest!: ball_conj_distrib [THEN iffD1] bchoice)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   537
apply (rule_tac x = "\<Union>A \<in> S. f A" in exI)
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   538
apply (auto intro: leadsTo_UN)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   539
(*Blast_tac says PROOF FAILED*)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   540
apply (rule_tac I1=S and A1="%i. f i - B" and A'1="%i. f i \<union> B" 
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   541
       in constrains_UN [THEN constrains_weaken], auto) 
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   542
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   543
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   544
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   545
text\<open>Misra's property W5\<close>
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   546
lemma wlt_constrains_wlt: "F \<in> (wlt F B - B) co (wlt F B)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   547
proof -
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   548
  from wlt_leadsTo [of F B, THEN leadsTo_123]
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   549
  show ?thesis
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   550
  proof (elim exE conjE)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   551
(* assumes have to be in exactly the form as in the goal displayed at
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   552
   this point.  Isar doesn't give you any automation. *)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   553
    fix C
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   554
    assume wlt: "wlt F B \<subseteq> C"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   555
       and lt:  "F \<in> C leadsTo B"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   556
       and co:  "F \<in> C - B co C \<union> B"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   557
    have eq: "C = wlt F B"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   558
    proof -
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   559
      from lt and wlt show ?thesis 
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   560
           by (blast dest: leadsTo_eq_subset_wlt [THEN iffD1])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   561
    qed
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   562
    from co show ?thesis by (simp add: eq wlt_increasing Un_absorb2)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   563
  qed
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13797
diff changeset
   564
qed
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   565
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   566
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   567
subsection\<open>Completion: Binary and General Finite versions\<close>
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   568
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   569
lemma completion_lemma :
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   570
     "[| W = wlt F (B' \<union> C);      
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   571
       F \<in> A leadsTo (A' \<union> C);  F \<in> A' co (A' \<union> C);    
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   572
       F \<in> B leadsTo (B' \<union> C);  F \<in> B' co (B' \<union> C) |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   573
    ==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B') \<union> C)"
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   574
apply (subgoal_tac "F \<in> (W-C) co (W \<union> B' \<union> C) ")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   575
 prefer 2
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   576
 apply (blast intro: wlt_constrains_wlt [THEN [2] constrains_Un, 
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   577
                                         THEN constrains_weaken])
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   578
apply (subgoal_tac "F \<in> (W-C) co W")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   579
 prefer 2
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   580
 apply (simp add: wlt_increasing Un_assoc Un_absorb2)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   581
apply (subgoal_tac "F \<in> (A \<inter> W - C) leadsTo (A' \<inter> W \<union> C) ")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   582
 prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   583
(** LEVEL 6 **)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   584
apply (subgoal_tac "F \<in> (A' \<inter> W \<union> C) leadsTo (A' \<inter> B' \<union> C) ")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   585
 prefer 2
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   586
 apply (rule leadsTo_Un_duplicate2)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   587
 apply (blast intro: leadsTo_Un_Un wlt_leadsTo
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   588
                         [THEN psp2, THEN leadsTo_weaken] leadsTo_refl)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   589
apply (drule leadsTo_Diff)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   590
apply (blast intro: subset_imp_leadsTo)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   591
apply (subgoal_tac "A \<inter> B \<subseteq> A \<inter> W")
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   592
 prefer 2
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   593
 apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   594
apply (blast intro: leadsTo_Trans subset_imp_leadsTo)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   595
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   596
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   597
lemmas completion = completion_lemma [OF refl]
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   598
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   599
lemma finite_completion_lemma:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   600
     "finite I ==> (\<forall>i \<in> I. F \<in> (A i) leadsTo (A' i \<union> C)) -->   
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   601
                   (\<forall>i \<in> I. F \<in> (A' i) co (A' i \<union> C)) -->  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   602
                   F \<in> (\<Inter>i \<in> I. A i) leadsTo ((\<Inter>i \<in> I. A' i) \<union> C)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   603
apply (erule finite_induct, auto)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   604
apply (rule completion)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   605
   prefer 4
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   606
   apply (simp only: INT_simps [symmetric])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   607
   apply (rule constrains_INT, auto)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   608
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   609
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   610
lemma finite_completion: 
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   611
     "[| finite I;   
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   612
         !!i. i \<in> I ==> F \<in> (A i) leadsTo (A' i \<union> C);  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   613
         !!i. i \<in> I ==> F \<in> (A' i) co (A' i \<union> C) |]    
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   614
      ==> F \<in> (\<Inter>i \<in> I. A i) leadsTo ((\<Inter>i \<in> I. A' i) \<union> C)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   615
by (blast intro: finite_completion_lemma [THEN mp, THEN mp])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   616
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   617
lemma stable_completion: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   618
     "[| F \<in> A leadsTo A';  F \<in> stable A';    
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   619
         F \<in> B leadsTo B';  F \<in> stable B' |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   620
    ==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B')"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   621
apply (unfold stable_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   622
apply (rule_tac C1 = "{}" in completion [THEN leadsTo_weaken_R])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   623
apply (force+)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   624
done
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   625
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   626
lemma finite_stable_completion: 
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   627
     "[| finite I;   
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   628
         !!i. i \<in> I ==> F \<in> (A i) leadsTo (A' i);  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   629
         !!i. i \<in> I ==> F \<in> stable (A' i) |]    
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   630
      ==> F \<in> (\<Inter>i \<in> I. A i) leadsTo (\<Inter>i \<in> I. A' i)"
13797
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   631
apply (unfold stable_def)
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   632
apply (rule_tac C1 = "{}" in finite_completion [THEN leadsTo_weaken_R])
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   633
apply (simp_all (no_asm_simp))
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   634
apply blast+
baefae13ad37 conversion of UNITY theories to new-style
paulson
parents: 10834
diff changeset
   635
done
9685
6d123a7e30bd xsymbols for leads-to and Join
paulson
parents: 8006
diff changeset
   636
35422
e74b6f3b950c tuned final whitespace;
wenzelm
parents: 35417
diff changeset
   637
end