src/HOL/UNITY/ProgressSets.thy
author immler
Sun, 03 Nov 2019 21:46:46 -0500
changeset 71034 e0755162093f
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
replace approximation oracle by less ad-hoc @{computation}s
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
     1
(*  Title:      HOL/UNITY/ProgressSets.thy
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     3
    Copyright   2003  University of Cambridge
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     4
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     5
Progress Sets.  From 
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     6
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     7
    David Meier and Beverly Sanders,
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     8
    Composing Leads-to Properties
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
     9
    Theoretical Computer Science 243:1-2 (2000), 339-361.
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    10
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    11
    David Meier,
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    12
    Progress Properties in Program Refinement and Parallel Composition
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    13
    Swiss Federal Institute of Technology Zurich (1997)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    14
*)
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    15
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    16
section\<open>Progress Sets\<close>
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    17
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15102
diff changeset
    18
theory ProgressSets imports Transformers begin
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    19
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
    20
subsection \<open>Complete Lattices and the Operator \<^term>\<open>cl\<close>\<close>
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
    21
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    22
definition lattice :: "'a set set => bool" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
    23
   \<comment> \<open>Meier calls them closure sets, but they are just complete lattices\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    24
   "lattice L ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
    25
         (\<forall>M. M \<subseteq> L --> \<Inter>M \<in> L) & (\<forall>M. M \<subseteq> L --> \<Union>M \<in> L)"
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    26
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    27
definition cl :: "['a set set, 'a set] => 'a set" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
    28
   \<comment> \<open>short for ``closure''\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    29
   "cl L r == \<Inter>{x. x\<in>L & r \<subseteq> x}"
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    30
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    31
lemma UNIV_in_lattice: "lattice L ==> UNIV \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    32
by (force simp add: lattice_def)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    33
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    34
lemma empty_in_lattice: "lattice L ==> {} \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    35
by (force simp add: lattice_def)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    36
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    37
lemma Union_in_lattice: "[|M \<subseteq> L; lattice L|] ==> \<Union>M \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    38
by (simp add: lattice_def)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    39
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    40
lemma Inter_in_lattice: "[|M \<subseteq> L; lattice L|] ==> \<Inter>M \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    41
by (simp add: lattice_def)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    42
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    43
lemma UN_in_lattice:
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    44
     "[|lattice L; !!i. i\<in>I ==> r i \<in> L|] ==> (\<Union>i\<in>I. r i) \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    45
apply (blast intro: Union_in_lattice) 
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    46
done
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    47
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    48
lemma INT_in_lattice:
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    49
     "[|lattice L; !!i. i\<in>I ==> r i \<in> L|] ==> (\<Inter>i\<in>I. r i)  \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    50
apply (blast intro: Inter_in_lattice) 
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    51
done
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    52
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    53
lemma Un_in_lattice: "[|x\<in>L; y\<in>L; lattice L|] ==> x\<union>y \<in> L"
44106
0e018cbcc0de tuned proofs
haftmann
parents: 35416
diff changeset
    54
  using Union_in_lattice [of "{x, y}" L] by simp
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    55
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    56
lemma Int_in_lattice: "[|x\<in>L; y\<in>L; lattice L|] ==> x\<inter>y \<in> L"
44106
0e018cbcc0de tuned proofs
haftmann
parents: 35416
diff changeset
    57
  using Inter_in_lattice [of "{x, y}" L] by simp
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    58
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    59
lemma lattice_stable: "lattice {X. F \<in> stable X}"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    60
by (simp add: lattice_def stable_def constrains_def, blast)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    61
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
    62
text\<open>The next three results state that \<^term>\<open>cl L r\<close> is the minimal
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
    63
 element of \<^term>\<open>L\<close> that includes \<^term>\<open>r\<close>.\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    64
lemma cl_in_lattice: "lattice L ==> cl L r \<in> L"
69144
f13b82281715 new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    65
  by (simp add: lattice_def cl_def)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    66
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    67
lemma cl_least: "[|c\<in>L; r\<subseteq>c|] ==> cl L r \<subseteq> c" 
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    68
by (force simp add: cl_def)
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    69
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    70
text\<open>The next three lemmas constitute assertion (4.61)\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    71
lemma cl_mono: "r \<subseteq> r' ==> cl L r \<subseteq> cl L r'"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    72
by (simp add: cl_def, blast)
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    73
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    74
lemma subset_cl: "r \<subseteq> cl L r"
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44106
diff changeset
    75
by (simp add: cl_def le_Inf_iff)
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    76
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    77
text\<open>A reformulation of @{thm subset_cl}\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
    78
lemma clI: "x \<in> r ==> x \<in> cl L r"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
    79
by (simp add: cl_def, blast)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
    80
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    81
text\<open>A reformulation of @{thm cl_least}\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
    82
lemma clD: "[|c \<in> cl L r; B \<in> L; r \<subseteq> B|] ==> c \<in> B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
    83
by (force simp add: cl_def)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
    84
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    85
lemma cl_UN_subset: "(\<Union>i\<in>I. cl L (r i)) \<subseteq> cl L (\<Union>i\<in>I. r i)"
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    86
by (simp add: cl_def, blast)
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    87
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    88
lemma cl_Un: "lattice L ==> cl L (r\<union>s) = cl L r \<union> cl L s"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    89
apply (rule equalityI) 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    90
 prefer 2 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    91
  apply (simp add: cl_def, blast)
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    92
apply (rule cl_least)
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    93
 apply (blast intro: Un_in_lattice cl_in_lattice)
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    94
apply (blast intro: subset_cl [THEN subsetD])  
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    95
done
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    96
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
    97
lemma cl_UN: "lattice L ==> cl L (\<Union>i\<in>I. r i) = (\<Union>i\<in>I. cl L (r i))"
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
    98
apply (rule equalityI) 
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
    99
 prefer 2 apply (simp add: cl_def, blast)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   100
apply (rule cl_least)
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   101
 apply (blast intro: UN_in_lattice cl_in_lattice)
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   102
apply (blast intro: subset_cl [THEN subsetD])  
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   103
done
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   104
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   105
lemma cl_Int_subset: "cl L (r\<inter>s) \<subseteq> cl L r \<inter> cl L s"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   106
by (simp add: cl_def, blast)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   107
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   108
lemma cl_idem [simp]: "cl L (cl L r) = cl L r"
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   109
by (simp add: cl_def, blast)
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   110
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   111
lemma cl_ident: "r\<in>L ==> cl L r = r" 
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   112
by (force simp add: cl_def)
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   113
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   114
lemma cl_empty [simp]: "lattice L ==> cl L {} = {}"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   115
by (simp add: cl_ident empty_in_lattice)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   116
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   117
lemma cl_UNIV [simp]: "lattice L ==> cl L UNIV = UNIV"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   118
by (simp add: cl_ident UNIV_in_lattice)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   119
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   120
text\<open>Assertion (4.62)\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   121
lemma cl_ident_iff: "lattice L ==> (cl L r = r) = (r\<in>L)" 
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   122
apply (rule iffI) 
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   123
 apply (erule subst)
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   124
 apply (erule cl_in_lattice)  
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   125
apply (erule cl_ident) 
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   126
done
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   127
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   128
lemma cl_subset_in_lattice: "[|cl L r \<subseteq> r; lattice L|] ==> r\<in>L" 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   129
by (simp add: cl_ident_iff [symmetric] equalityI subset_cl)
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   130
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   131
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   132
subsection \<open>Progress Sets and the Main Lemma\<close>
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   133
text\<open>A progress set satisfies certain closure conditions and is a 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   134
simple way of including the set \<^term>\<open>wens_set F B\<close>.\<close>
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   135
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   136
definition closed :: "['a program, 'a set, 'a set,  'a set set] => bool" where
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   137
   "closed F T B L == \<forall>M. \<forall>act \<in> Acts F. B\<subseteq>M & T\<inter>M \<in> L -->
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   138
                              T \<inter> (B \<union> wp act M) \<in> L"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   139
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   140
definition progress_set :: "['a program, 'a set, 'a set] => 'a set set set" where
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   141
   "progress_set F T B ==
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   142
      {L. lattice L & B \<in> L & T \<in> L & closed F T B L}"
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   143
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   144
lemma closedD:
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   145
   "[|closed F T B L; act \<in> Acts F; B\<subseteq>M; T\<inter>M \<in> L|] 
14150
9a23e4eb5eb3 A document for UNITY
paulson
parents: 13888
diff changeset
   146
    ==> T \<inter> (B \<union> wp act M) \<in> L" 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   147
by (simp add: closed_def) 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   148
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   149
text\<open>Note: the formalization below replaces Meier's \<^term>\<open>q\<close> by \<^term>\<open>B\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   150
and \<^term>\<open>m\<close> by \<^term>\<open>X\<close>.\<close>
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   151
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   152
text\<open>Part of the proof of the claim at the bottom of page 97.  It's
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   153
proved separately because the argument requires a generalization over
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   154
all \<^term>\<open>act \<in> Acts F\<close>.\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   155
lemma lattice_awp_lemma:
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
   156
  assumes TXC:  "T\<inter>X \<in> C" \<comment> \<open>induction hypothesis in theorem below\<close>
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
   157
      and BsubX:  "B \<subseteq> X"   \<comment> \<open>holds in inductive step\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   158
      and latt: "lattice C"
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   159
      and TC:   "T \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   160
      and BC:   "B \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   161
      and clos: "closed F T B C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   162
    shows "T \<inter> (B \<union> awp F (X \<union> cl C (T\<inter>r))) \<in> C"
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   163
apply (simp del: INT_simps add: awp_def INT_extend_simps) 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   164
apply (rule INT_in_lattice [OF latt]) 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   165
apply (erule closedD [OF clos]) 
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   166
apply (simp add: subset_trans [OF BsubX Un_upper1]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   167
apply (subgoal_tac "T \<inter> (X \<union> cl C (T\<inter>r)) = (T\<inter>X) \<union> cl C (T\<inter>r)")
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   168
 prefer 2 apply (blast intro: TC clD) 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   169
apply (erule ssubst) 
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   170
apply (blast intro: Un_in_lattice latt cl_in_lattice TXC) 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   171
done
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   172
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   173
text\<open>Remainder of the proof of the claim at the bottom of page 97.\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   174
lemma lattice_lemma:
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
   175
  assumes TXC:  "T\<inter>X \<in> C" \<comment> \<open>induction hypothesis in theorem below\<close>
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
   176
      and BsubX:  "B \<subseteq> X"   \<comment> \<open>holds in inductive step\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   177
      and act:  "act \<in> Acts F"
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   178
      and latt: "lattice C"
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   179
      and TC:   "T \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   180
      and BC:   "B \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   181
      and clos: "closed F T B C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   182
    shows "T \<inter> (wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)) \<union> X) \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   183
apply (subgoal_tac "T \<inter> (B \<union> wp act X) \<in> C")
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   184
 prefer 2 apply (simp add: closedD [OF clos] act BsubX TXC)
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   185
apply (drule Int_in_lattice
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   186
              [OF _ lattice_awp_lemma [OF TXC BsubX latt TC BC clos, of r]
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   187
                    latt])
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   188
apply (subgoal_tac
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
   189
         "T \<inter> (B \<union> wp act X) \<inter> (T \<inter> (B \<union> awp F (X \<union> cl C (T\<inter>r)))) = 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
   190
          T \<inter> (B \<union> wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)))") 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   191
 prefer 2 apply blast 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   192
apply simp  
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   193
apply (drule Un_in_lattice [OF _ TXC latt])  
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   194
apply (subgoal_tac
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
   195
         "T \<inter> (B \<union> wp act X \<inter> awp F (X \<union> cl C (T\<inter>r))) \<union> T\<inter>X = 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
   196
          T \<inter> (wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)) \<union> X)")
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   197
 apply simp 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   198
apply (blast intro: BsubX [THEN subsetD]) 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   199
done
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   200
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   201
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   202
text\<open>Induction step for the main lemma\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   203
lemma progress_induction_step:
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
   204
  assumes TXC:  "T\<inter>X \<in> C" \<comment> \<open>induction hypothesis in theorem below\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   205
      and act:  "act \<in> Acts F"
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   206
      and Xwens: "X \<in> wens_set F B"
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   207
      and latt: "lattice C"
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   208
      and  TC:  "T \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   209
      and  BC:  "B \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   210
      and clos: "closed F T B C"
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   211
      and Fstable: "F \<in> stable T"
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   212
  shows "T \<inter> wens F act X \<in> C"
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   213
proof -
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   214
  from Xwens have BsubX: "B \<subseteq> X"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   215
    by (rule wens_set_imp_subset) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   216
  let ?r = "wens F act X"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   217
  have "?r \<subseteq> (wp act X \<inter> awp F (X\<union>?r)) \<union> X"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   218
    by (simp add: wens_unfold [symmetric])
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   219
  then have "T\<inter>?r \<subseteq> T \<inter> ((wp act X \<inter> awp F (X\<union>?r)) \<union> X)"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   220
    by blast
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   221
  then have "T\<inter>?r \<subseteq> T \<inter> ((wp act X \<inter> awp F (T \<inter> (X\<union>?r))) \<union> X)"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   222
    by (simp add: awp_Int_eq Fstable stable_imp_awp_ident, blast) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   223
  then have "T\<inter>?r \<subseteq> T \<inter> ((wp act X \<inter> awp F (X \<union> cl C (T\<inter>?r))) \<union> X)"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   224
    by (blast intro: awp_mono [THEN [2] rev_subsetD] subset_cl [THEN subsetD])
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   225
  then have "cl C (T\<inter>?r) \<subseteq> 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   226
             cl C (T \<inter> ((wp act X \<inter> awp F (X \<union> cl C (T\<inter>?r))) \<union> X))"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   227
    by (rule cl_mono) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   228
  then have "cl C (T\<inter>?r) \<subseteq> 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   229
             T \<inter> ((wp act X \<inter> awp F (X \<union> cl C (T\<inter>?r))) \<union> X)"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   230
    by (simp add: cl_ident lattice_lemma [OF TXC BsubX act latt TC BC clos])
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   231
  then have "cl C (T\<inter>?r) \<subseteq> (wp act X \<inter> awp F (X \<union> cl C (T\<inter>?r))) \<union> X"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   232
    by blast
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   233
  then have "cl C (T\<inter>?r) \<subseteq> ?r"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   234
    by (blast intro!: subset_wens) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   235
  then have cl_subset: "cl C (T\<inter>?r) \<subseteq> T\<inter>?r"
32693
6c6b1ba5e71e tuned proofs
haftmann
parents: 32604
diff changeset
   236
    by (simp add: cl_ident TC
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   237
                  subset_trans [OF cl_mono [OF Int_lower1]]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   238
  show ?thesis
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   239
    by (rule cl_subset_in_lattice [OF cl_subset latt]) 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   240
qed
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   241
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   242
text\<open>Proved on page 96 of Meier's thesis.  The special case when
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   243
   \<^term>\<open>T=UNIV\<close> states that every progress set for the program \<^term>\<open>F\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   244
   and set \<^term>\<open>B\<close> includes the set \<^term>\<open>wens_set F B\<close>.\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   245
lemma progress_set_lemma:
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   246
     "[|C \<in> progress_set F T B; r \<in> wens_set F B; F \<in> stable T|] ==> T\<inter>r \<in> C"
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   247
apply (simp add: progress_set_def, clarify) 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   248
apply (erule wens_set.induct) 
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   249
  txt\<open>Base\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   250
  apply (simp add: Int_in_lattice) 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   251
 txt\<open>The difficult \<^term>\<open>wens\<close> case\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   252
 apply (simp add: progress_induction_step) 
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   253
txt\<open>Disjunctive case\<close>
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   254
apply (subgoal_tac "(\<Union>U\<in>W. T \<inter> U) \<in> C") 
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 51488
diff changeset
   255
 apply simp 
13861
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   256
apply (blast intro: UN_in_lattice) 
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   257
done
0c18f31d901a Proved the main lemma on progress sets
paulson
parents: 13853
diff changeset
   258
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   259
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   260
subsection \<open>The Progress Set Union Theorem\<close>
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   261
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   262
lemma closed_mono:
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   263
  assumes BB':  "B \<subseteq> B'"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   264
      and TBwp: "T \<inter> (B \<union> wp act M) \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   265
      and B'C:  "B' \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   266
      and TC:   "T \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   267
      and latt: "lattice C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   268
  shows "T \<inter> (B' \<union> wp act M) \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   269
proof -
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   270
  from TBwp have "(T\<inter>B) \<union> (T \<inter> wp act M) \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   271
    by (simp add: Int_Un_distrib)
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   272
  then have TBBC: "(T\<inter>B') \<union> ((T\<inter>B) \<union> (T \<inter> wp act M)) \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   273
    by (blast intro: Int_in_lattice Un_in_lattice TC B'C latt) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   274
  show ?thesis
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   275
    by (rule eqelem_imp_iff [THEN iffD1, OF _ TBBC], 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   276
        blast intro: BB' [THEN subsetD]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   277
qed
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   278
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   279
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   280
lemma progress_set_mono:
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   281
    assumes BB':  "B \<subseteq> B'"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   282
    shows
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   283
     "[| B' \<in> C;  C \<in> progress_set F T B|] 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   284
      ==> C \<in> progress_set F T B'"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   285
by (simp add: progress_set_def closed_def closed_mono [OF BB'] 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   286
                 subset_trans [OF BB']) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   287
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   288
theorem progress_set_Union:
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   289
  assumes leadsTo: "F \<in> A leadsTo B'"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   290
      and prog: "C \<in> progress_set F T B"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   291
      and Fstable: "F \<in> stable T"
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   292
      and BB':  "B \<subseteq> B'"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   293
      and B'C:  "B' \<in> C"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   294
      and Gco: "!!X. X\<in>C ==> G \<in> X-B co X"
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   295
  shows "F\<squnion>G \<in> T\<inter>A leadsTo B'"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   296
apply (insert prog Fstable) 
13866
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   297
apply (rule leadsTo_Join [OF leadsTo]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   298
  apply (force simp add: progress_set_def awp_iff_stable [symmetric]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   299
apply (simp add: awp_iff_constrains)
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   300
apply (drule progress_set_mono [OF BB' B'C]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   301
apply (blast intro: progress_set_lemma Gco constrains_weaken_L 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   302
                    BB' [THEN subsetD]) 
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   303
done
b42d7983a822 More "progress set" material
paulson
parents: 13861
diff changeset
   304
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   305
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   306
subsection \<open>Some Progress Sets\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   307
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   308
lemma UNIV_in_progress_set: "UNIV \<in> progress_set F T B"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   309
by (simp add: progress_set_def lattice_def closed_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents: 13866
diff changeset
   310
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   311
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   312
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   313
subsubsection \<open>Lattices and Relations\<close>
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   314
text\<open>From Meier's thesis, section 4.5.3\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   315
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   316
definition relcl :: "'a set set => ('a * 'a) set" where
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   317
    \<comment> \<open>Derived relation from a lattice\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   318
    "relcl L == {(x,y). y \<in> cl L {x}}"
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   319
  
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   320
definition latticeof :: "('a * 'a) set => 'a set set" where
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   321
    \<comment> \<open>Derived lattice from a relation: the set of upwards-closed sets\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   322
    "latticeof r == {X. \<forall>s t. s \<in> X & (s,t) \<in> r --> t \<in> X}"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   323
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   324
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   325
lemma relcl_refl: "(a,a) \<in> relcl L"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   326
by (simp add: relcl_def subset_cl [THEN subsetD])
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   327
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   328
lemma relcl_trans:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   329
     "[| (a,b) \<in> relcl L; (b,c) \<in> relcl L; lattice L |] ==> (a,c) \<in> relcl L"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   330
apply (simp add: relcl_def)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   331
apply (blast intro: clD cl_in_lattice)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   332
done
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   333
30198
922f944f03b2 name changes
nipkow
parents: 23767
diff changeset
   334
lemma refl_relcl: "lattice L ==> refl (relcl L)"
922f944f03b2 name changes
nipkow
parents: 23767
diff changeset
   335
by (simp add: refl_onI relcl_def subset_cl [THEN subsetD])
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   336
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   337
lemma trans_relcl: "lattice L ==> trans (relcl L)"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   338
by (blast intro: relcl_trans transI)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   339
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   340
lemma lattice_latticeof: "lattice (latticeof r)"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   341
by (auto simp add: lattice_def latticeof_def)
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   342
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   343
lemma lattice_singletonI:
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   344
     "[|lattice L; !!s. s \<in> X ==> {s} \<in> L|] ==> X \<in> L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   345
apply (cut_tac UN_singleton [of X]) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   346
apply (erule subst) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   347
apply (simp only: UN_in_lattice) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   348
done
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   349
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   350
text\<open>Equation (4.71) of Meier's thesis.  He gives no proof.\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   351
lemma cl_latticeof:
30198
922f944f03b2 name changes
nipkow
parents: 23767
diff changeset
   352
     "[|refl r; trans r|] 
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   353
      ==> cl (latticeof r) X = {t. \<exists>s. s\<in>X & (s,t) \<in> r}" 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   354
apply (rule equalityI) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   355
 apply (rule cl_least) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   356
  apply (simp (no_asm_use) add: latticeof_def trans_def, blast)
30198
922f944f03b2 name changes
nipkow
parents: 23767
diff changeset
   357
 apply (simp add: latticeof_def refl_on_def, blast)
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   358
apply (simp add: latticeof_def, clarify)
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   359
apply (unfold cl_def, blast) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   360
done
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   361
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   362
text\<open>Related to (4.71).\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   363
lemma cl_eq_Collect_relcl:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   364
     "lattice L ==> cl L X = {t. \<exists>s. s\<in>X & (s,t) \<in> relcl L}" 
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   365
apply (cut_tac UN_singleton [of X]) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   366
apply (erule subst) 
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   367
apply (force simp only: relcl_def cl_UN)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   368
done
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   369
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   370
text\<open>Meier's theorem of section 4.5.3\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   371
theorem latticeof_relcl_eq: "lattice L ==> latticeof (relcl L) = L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   372
apply (rule equalityI) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   373
 prefer 2 apply (force simp add: latticeof_def relcl_def cl_def, clarify) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   374
apply (rename_tac X)
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   375
apply (rule cl_subset_in_lattice)   
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   376
 prefer 2 apply assumption
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   377
apply (drule cl_ident_iff [OF lattice_latticeof, THEN iffD2])
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   378
apply (drule equalityD1)   
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   379
apply (rule subset_trans) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   380
 prefer 2 apply assumption
59807
22bc39064290 prefer local fixes;
wenzelm
parents: 58889
diff changeset
   381
apply (thin_tac "_ \<subseteq> X") 
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   382
apply (cut_tac A=X in UN_singleton) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   383
apply (erule subst) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   384
apply (simp only: cl_UN lattice_latticeof 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   385
                  cl_latticeof [OF refl_relcl trans_relcl]) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   386
apply (simp add: relcl_def) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   387
done
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   388
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   389
theorem relcl_latticeof_eq:
30198
922f944f03b2 name changes
nipkow
parents: 23767
diff changeset
   390
     "[|refl r; trans r|] ==> relcl (latticeof r) = r"
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   391
by (simp add: relcl_def cl_latticeof)
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   392
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   393
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   394
subsubsection \<open>Decoupling Theorems\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   395
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   396
definition decoupled :: "['a program, 'a program] => bool" where
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   397
   "decoupled F G ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32693
diff changeset
   398
        \<forall>act \<in> Acts F. \<forall>B. G \<in> stable B --> G \<in> stable (wp act B)"
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   399
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   400
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   401
text\<open>Rao's Decoupling Theorem\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   402
lemma stableco: "F \<in> stable A ==> F \<in> A-B co A"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   403
by (simp add: stable_def constrains_def, blast) 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   404
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   405
theorem decoupling:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   406
  assumes leadsTo: "F \<in> A leadsTo B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   407
      and Gstable: "G \<in> stable B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   408
      and dec:     "decoupled F G"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   409
  shows "F\<squnion>G \<in> A leadsTo B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   410
proof -
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   411
  have prog: "{X. G \<in> stable X} \<in> progress_set F UNIV B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   412
    by (simp add: progress_set_def lattice_stable Gstable closed_def
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   413
                  stable_Un [OF Gstable] dec [unfolded decoupled_def]) 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   414
  have "F\<squnion>G \<in> (UNIV\<inter>A) leadsTo B" 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   415
    by (rule progress_set_Union [OF leadsTo prog],
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   416
        simp_all add: Gstable stableco)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   417
  thus ?thesis by simp
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   418
qed
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   419
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   420
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   421
text\<open>Rao's Weak Decoupling Theorem\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   422
theorem weak_decoupling:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   423
  assumes leadsTo: "F \<in> A leadsTo B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   424
      and stable: "F\<squnion>G \<in> stable B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   425
      and dec:     "decoupled F (F\<squnion>G)"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   426
  shows "F\<squnion>G \<in> A leadsTo B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   427
proof -
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   428
  have prog: "{X. F\<squnion>G \<in> stable X} \<in> progress_set F UNIV B" 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   429
    by (simp del: Join_stable
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   430
             add: progress_set_def lattice_stable stable closed_def
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   431
                  stable_Un [OF stable] dec [unfolded decoupled_def])
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   432
  have "F\<squnion>G \<in> (UNIV\<inter>A) leadsTo B" 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   433
    by (rule progress_set_Union [OF leadsTo prog],
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   434
        simp_all del: Join_stable add: stable,
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   435
        simp add: stableco) 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   436
  thus ?thesis by simp
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   437
qed
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   438
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   439
text\<open>The ``Decoupling via \<^term>\<open>G'\<close> Union Theorem''\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   440
theorem decoupling_via_aux:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   441
  assumes leadsTo: "F \<in> A leadsTo B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   442
      and prog: "{X. G' \<in> stable X} \<in> progress_set F UNIV B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   443
      and GG':  "G \<le> G'"  
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63146
diff changeset
   444
               \<comment> \<open>Beware!  This is the converse of the refinement relation!\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   445
  shows "F\<squnion>G \<in> A leadsTo B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   446
proof -
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   447
  from prog have stable: "G' \<in> stable B"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   448
    by (simp add: progress_set_def)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   449
  have "F\<squnion>G \<in> (UNIV\<inter>A) leadsTo B" 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   450
    by (rule progress_set_Union [OF leadsTo prog],
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   451
        simp_all add: stable stableco component_stable [OF GG'])
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   452
  thus ?thesis by simp
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   453
qed
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   454
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   455
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   456
subsection\<open>Composition Theorems Based on Monotonicity and Commutativity\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   457
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   458
subsubsection\<open>Commutativity of \<^term>\<open>cl L\<close> and assignment.\<close>
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   459
definition commutes :: "['a program, 'a set, 'a set,  'a set set] => bool" where
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   460
   "commutes F T B L ==
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   461
       \<forall>M. \<forall>act \<in> Acts F. B \<subseteq> M --> 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   462
           cl L (T \<inter> wp act M) \<subseteq> T \<inter> (B \<union> wp act (cl L (T\<inter>M)))"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   463
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   464
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   465
text\<open>From Meier's thesis, section 4.5.6\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   466
lemma commutativity1_lemma:
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   467
  assumes commutes: "commutes F T B L" 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   468
      and lattice:  "lattice L"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   469
      and BL: "B \<in> L"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   470
      and TL: "T \<in> L"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   471
  shows "closed F T B L"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   472
apply (simp add: closed_def, clarify)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   473
apply (rule ProgressSets.cl_subset_in_lattice [OF _ lattice])  
32693
6c6b1ba5e71e tuned proofs
haftmann
parents: 32604
diff changeset
   474
apply (simp add: Int_Un_distrib cl_Un [OF lattice] 
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   475
                 cl_ident Int_in_lattice [OF TL BL lattice] Un_upper1)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   476
apply (subgoal_tac "cl L (T \<inter> wp act M) \<subseteq> T \<inter> (B \<union> wp act (cl L (T \<inter> M)))") 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   477
 prefer 2 
15102
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 14150
diff changeset
   478
 apply (cut_tac commutes, simp add: commutes_def) 
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   479
apply (erule subset_trans) 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   480
apply (simp add: cl_ident)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   481
apply (blast intro: rev_subsetD [OF _ wp_mono]) 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   482
done
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   483
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   484
text\<open>Version packaged with @{thm progress_set_Union}\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   485
lemma commutativity1:
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   486
  assumes leadsTo: "F \<in> A leadsTo B"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   487
      and lattice:  "lattice L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   488
      and BL: "B \<in> L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   489
      and TL: "T \<in> L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   490
      and Fstable: "F \<in> stable T"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   491
      and Gco: "!!X. X\<in>L ==> G \<in> X-B co X"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   492
      and commutes: "commutes F T B L" 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   493
  shows "F\<squnion>G \<in> T\<inter>A leadsTo B"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   494
by (rule progress_set_Union [OF leadsTo _ Fstable subset_refl BL Gco],
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   495
    simp add: progress_set_def commutativity1_lemma commutes lattice BL TL) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   496
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   497
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   498
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69144
diff changeset
   499
text\<open>Possibly move to Relation.thy, after \<^term>\<open>single_valued\<close>\<close>
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   500
definition funof :: "[('a*'b)set, 'a] => 'b" where
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   501
   "funof r == (\<lambda>x. THE y. (x,y) \<in> r)"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   502
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   503
lemma funof_eq: "[|single_valued r; (x,y) \<in> r|] ==> funof r x = y"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   504
by (simp add: funof_def single_valued_def, blast)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   505
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   506
lemma funof_Pair_in:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   507
     "[|single_valued r; x \<in> Domain r|] ==> (x, funof r x) \<in> r"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   508
by (force simp add: funof_eq) 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   509
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   510
lemma funof_in:
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   511
     "[|r``{x} \<subseteq> A; single_valued r; x \<in> Domain r|] ==> funof r x \<in> A" 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   512
by (force simp add: funof_eq)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   513
 
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   514
lemma funof_imp_wp: "[|funof act t \<in> A; single_valued act|] ==> t \<in> wp act A"
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   515
by (force simp add: in_wp_iff funof_eq)
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   516
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   517
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   518
subsubsection\<open>Commutativity of Functions and Relation\<close>
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   519
text\<open>Thesis, page 109\<close>
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   520
32604
8b3e2bc91a46 partially isarified proof
haftmann
parents: 32139
diff changeset
   521
(*FIXME: this proof is still an ungodly mess*)
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   522
text\<open>From Meier's thesis, section 4.5.6\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   523
lemma commutativity2_lemma:
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   524
  assumes dcommutes: 
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   525
      "\<And>act s t. act \<in> Acts F \<Longrightarrow> s \<in> T \<Longrightarrow> (s, t) \<in> relcl L \<Longrightarrow>
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   526
        s \<in> B | t \<in> B | (funof act s, funof act t) \<in> relcl L"
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   527
    and determ: "!!act. act \<in> Acts F ==> single_valued act"
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   528
    and total: "!!act. act \<in> Acts F ==> Domain act = UNIV"
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   529
    and lattice:  "lattice L"
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   530
    and BL: "B \<in> L"
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   531
    and TL: "T \<in> L"
11d9c2768729 tuned proofs;
wenzelm
parents: 44918
diff changeset
   532
    and Fstable: "F \<in> stable T"
13874
0da2141606c6 More on progress sets
paulson
parents: 13870
diff changeset
   533
  shows  "commutes F T B L"
32604
8b3e2bc91a46 partially isarified proof
haftmann
parents: 32139
diff changeset
   534
proof -
51488
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   535
  { fix M and act and t
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   536
    assume 1: "B \<subseteq> M" "act \<in> Acts F" "t \<in> cl L (T \<inter> wp act M)"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   537
    then have "\<exists>s. (s,t) \<in> relcl L \<and> s \<in> T \<inter> wp act M"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   538
      by (force simp add: cl_eq_Collect_relcl [OF lattice])
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   539
    then obtain s where 2: "(s, t) \<in> relcl L" "s \<in> T" "s \<in> wp act M"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   540
      by blast
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   541
    then have 3: "\<forall>u\<in>L. s \<in> u --> t \<in> u"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   542
      apply (intro ballI impI) 
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   543
      apply (subst cl_ident [symmetric], assumption)
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   544
      apply (simp add: relcl_def)  
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   545
      apply (blast intro: cl_mono [THEN [2] rev_subsetD])
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   546
      done
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   547
    with 1 2 Fstable have 4: "funof act s \<in> T\<inter>M"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   548
      by (force intro!: funof_in 
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   549
        simp add: wp_def stable_def constrains_def determ total)
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   550
    with 1 2 3 have 5: "s \<in> B | t \<in> B | (funof act s, funof act t) \<in> relcl L"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   551
      by (intro dcommutes) assumption+ 
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   552
    with 1 2 3 4 have "t \<in> B | funof act t \<in> cl L (T\<inter>M)"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   553
      by (simp add: relcl_def) (blast intro: BL cl_mono [THEN [2] rev_subsetD])  
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   554
    with 1 2 3 4 5 have "t \<in> B | t \<in> wp act (cl L (T\<inter>M))"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   555
      by (blast intro: funof_imp_wp determ) 
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   556
    with 2 3 have "t \<in> T \<and> (t \<in> B \<or> t \<in> wp act (cl L (T \<inter> M)))"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   557
      by (blast intro: TL cl_mono [THEN [2] rev_subsetD])
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   558
    then have"t \<in> T \<inter> (B \<union> wp act (cl L (T \<inter> M)))"
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   559
      by simp
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   560
  }
3c886fe611b8 tuned proof
haftmann
parents: 45477
diff changeset
   561
  then show "commutes F T B L" unfolding commutes_def by clarify
32604
8b3e2bc91a46 partially isarified proof
haftmann
parents: 32139
diff changeset
   562
qed
8b3e2bc91a46 partially isarified proof
haftmann
parents: 32139
diff changeset
   563
  
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   564
text\<open>Version packaged with @{thm progress_set_Union}\<close>
13885
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   565
lemma commutativity2:
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   566
  assumes leadsTo: "F \<in> A leadsTo B"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   567
      and dcommutes: 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   568
        "\<forall>act \<in> Acts F. 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   569
         \<forall>s \<in> T. \<forall>t. (s,t) \<in> relcl L --> 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   570
                      s \<in> B | t \<in> B | (funof act s, funof act t) \<in> relcl L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   571
      and determ: "!!act. act \<in> Acts F ==> single_valued act"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   572
      and total: "!!act. act \<in> Acts F ==> Domain act = UNIV"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   573
      and lattice:  "lattice L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   574
      and BL: "B \<in> L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   575
      and TL: "T \<in> L"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   576
      and Fstable: "F \<in> stable T"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   577
      and Gco: "!!X. X\<in>L ==> G \<in> X-B co X"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   578
  shows "F\<squnion>G \<in> T\<inter>A leadsTo B"
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   579
apply (rule commutativity1 [OF leadsTo lattice]) 
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   580
apply (simp_all add: Gco commutativity2_lemma dcommutes determ total
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   581
                     lattice BL TL Fstable)
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   582
done
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   583
de6fac7d5351 Proofs for section 4.5.3
paulson
parents: 13874
diff changeset
   584
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   585
subsection \<open>Monotonicity\<close>
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   586
text\<open>From Meier's thesis, section 4.5.7, page 110\<close>
13888
16f424af58a2 more comments and tweaks
paulson
parents: 13885
diff changeset
   587
(*to be continued?*)
16f424af58a2 more comments and tweaks
paulson
parents: 13885
diff changeset
   588
13853
89131afa9f01 New theory ProgressSets. Definition of closure sets
paulson
parents:
diff changeset
   589
end