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