src/HOL/UNITY/Project.thy
author haftmann
Mon, 01 Mar 2010 13:40:23 +0100
changeset 35416 d8d7d1b785af
parent 32960 69916a850301
child 45477 11d9c2768729
permissions -rw-r--r--
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 16417
diff changeset
     1
(*  Title:      HOL/UNITY/Project.thy
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     3
    Copyright   1999  University of Cambridge
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     4
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
     5
Projections of state sets (also of actions and programs).
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     6
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
     7
Inheritance of GUARANTEES properties under extension.
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     8
*)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     9
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    10
header{*Projections of State Sets*}
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    11
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13819
diff changeset
    12
theory Project imports Extend begin
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    13
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    14
definition projecting :: "['c program => 'c set, 'a*'b => 'c, 
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    15
                  'a program, 'c program set, 'a program set] => bool" where
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8055
diff changeset
    16
    "projecting C h F X' X ==
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    17
       \<forall>G. extend h F\<squnion>G \<in> X' --> F\<squnion>project h (C G) G \<in> X"
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    18
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    19
definition extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, 
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    20
                 'c program set, 'a program set] => bool" where
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8055
diff changeset
    21
    "extending C h F Y' Y ==
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    22
       \<forall>G. extend h F  ok G --> F\<squnion>project h (C G) G \<in> Y
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
    23
              --> extend h F\<squnion>G \<in> Y'"
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8055
diff changeset
    24
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    25
definition subset_closed :: "'a set set => bool" where
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    26
    "subset_closed U == \<forall>A \<in> U. Pow A \<subseteq> U"  
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    27
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    28
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    29
lemma (in Extend) project_extend_constrains_I:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    30
     "F \<in> A co B ==> project h C (extend h F) \<in> A co B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    31
apply (auto simp add: extend_act_def project_act_def constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    32
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    33
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    34
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    35
subsection{*Safety*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    36
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    37
(*used below to prove Join_project_ensures*)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    38
lemma (in Extend) project_unless [rule_format]:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    39
     "[| G \<in> stable C;  project h C G \<in> A unless B |]  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    40
      ==> G \<in> (C \<inter> extend_set h A) unless (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    41
apply (simp add: unless_def project_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    42
apply (blast dest: stable_constrains_Int intro: constrains_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    43
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    44
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    45
(*Generalizes project_constrains to the program F\<squnion>project h C G
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    46
  useful with guarantees reasoning*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    47
lemma (in Extend) Join_project_constrains:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    48
     "(F\<squnion>project h C G \<in> A co B)  =   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    49
        (extend h F\<squnion>G \<in> (C \<inter> extend_set h A) co (extend_set h B) &   
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    50
         F \<in> A co B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    51
apply (simp (no_asm) add: project_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    52
apply (blast intro: extend_constrains [THEN iffD2, THEN constrains_weaken] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    53
             dest: constrains_imp_subset)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    54
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    55
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    56
(*The condition is required to prove the left-to-right direction
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    57
  could weaken it to G \<in> (C \<inter> extend_set h A) co C*)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    58
lemma (in Extend) Join_project_stable: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    59
     "extend h F\<squnion>G \<in> stable C  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    60
      ==> (F\<squnion>project h C G \<in> stable A)  =   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    61
          (extend h F\<squnion>G \<in> stable (C \<inter> extend_set h A) &   
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
    62
           F \<in> stable A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    63
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    64
apply (simp only: Join_project_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    65
apply (blast intro: constrains_weaken dest: constrains_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    66
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    67
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    68
(*For using project_guarantees in particular cases*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    69
lemma (in Extend) project_constrains_I:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    70
     "extend h F\<squnion>G \<in> extend_set h A co extend_set h B  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    71
      ==> F\<squnion>project h C G \<in> A co B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    72
apply (simp add: project_constrains extend_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    73
apply (blast intro: constrains_weaken dest: constrains_imp_subset)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    74
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    75
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    76
lemma (in Extend) project_increasing_I: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    77
     "extend h F\<squnion>G \<in> increasing (func o f)  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    78
      ==> F\<squnion>project h C G \<in> increasing func"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    79
apply (unfold increasing_def stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    80
apply (simp del: Join_constrains
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    81
            add: project_constrains_I extend_set_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    82
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    83
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    84
lemma (in Extend) Join_project_increasing:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    85
     "(F\<squnion>project h UNIV G \<in> increasing func)  =   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    86
      (extend h F\<squnion>G \<in> increasing (func o f))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    87
apply (rule iffI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    88
apply (erule_tac [2] project_increasing_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    89
apply (simp del: Join_stable
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    90
            add: increasing_def Join_project_stable)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    91
apply (auto simp add: extend_set_eq_Collect extend_stable [THEN iffD1])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    92
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    93
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    94
(*The UNIV argument is essential*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    95
lemma (in Extend) project_constrains_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    96
     "F\<squnion>project h UNIV G \<in> A co B  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
    97
      ==> extend h F\<squnion>G \<in> extend_set h A co extend_set h B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    98
by (simp add: project_constrains extend_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
    99
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   100
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   101
subsection{*"projecting" and union/intersection (no converses)*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   102
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   103
lemma projecting_Int: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   104
     "[| projecting C h F XA' XA;  projecting C h F XB' XB |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   105
      ==> projecting C h F (XA' \<inter> XB') (XA \<inter> XB)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   106
by (unfold projecting_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   107
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   108
lemma projecting_Un: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   109
     "[| projecting C h F XA' XA;  projecting C h F XB' XB |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   110
      ==> projecting C h F (XA' \<union> XB') (XA \<union> XB)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   111
by (unfold projecting_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   112
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   113
lemma projecting_INT: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   114
     "[| !!i. i \<in> I ==> projecting C h F (X' i) (X i) |]  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   115
      ==> projecting C h F (\<Inter>i \<in> I. X' i) (\<Inter>i \<in> I. X i)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   116
by (unfold projecting_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   117
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   118
lemma projecting_UN: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   119
     "[| !!i. i \<in> I ==> projecting C h F (X' i) (X i) |]  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   120
      ==> projecting C h F (\<Union>i \<in> I. X' i) (\<Union>i \<in> I. X i)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   121
by (unfold projecting_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   122
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   123
lemma projecting_weaken: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   124
     "[| projecting C h F X' X;  U'<=X';  X \<subseteq> U |] ==> projecting C h F U' U"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   125
by (unfold projecting_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   126
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   127
lemma projecting_weaken_L: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   128
     "[| projecting C h F X' X;  U'<=X' |] ==> projecting C h F U' X"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   129
by (unfold projecting_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   130
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   131
lemma extending_Int: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   132
     "[| extending C h F YA' YA;  extending C h F YB' YB |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   133
      ==> extending C h F (YA' \<inter> YB') (YA \<inter> YB)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   134
by (unfold extending_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   135
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   136
lemma extending_Un: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   137
     "[| extending C h F YA' YA;  extending C h F YB' YB |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   138
      ==> extending C h F (YA' \<union> YB') (YA \<union> YB)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   139
by (unfold extending_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   140
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   141
lemma extending_INT: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   142
     "[| !!i. i \<in> I ==> extending C h F (Y' i) (Y i) |]  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   143
      ==> extending C h F (\<Inter>i \<in> I. Y' i) (\<Inter>i \<in> I. Y i)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   144
by (unfold extending_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   145
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   146
lemma extending_UN: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   147
     "[| !!i. i \<in> I ==> extending C h F (Y' i) (Y i) |]  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   148
      ==> extending C h F (\<Union>i \<in> I. Y' i) (\<Union>i \<in> I. Y i)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   149
by (unfold extending_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   150
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   151
lemma extending_weaken: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   152
     "[| extending C h F Y' Y;  Y'<=V';  V \<subseteq> Y |] ==> extending C h F V' V"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   153
by (unfold extending_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   154
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   155
lemma extending_weaken_L: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   156
     "[| extending C h F Y' Y;  Y'<=V' |] ==> extending C h F V' Y"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   157
by (unfold extending_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   158
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   159
lemma projecting_UNIV: "projecting C h F X' UNIV"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   160
by (simp add: projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   161
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   162
lemma (in Extend) projecting_constrains: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   163
     "projecting C h F (extend_set h A co extend_set h B) (A co B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   164
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   165
apply (blast intro: project_constrains_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   166
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   167
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   168
lemma (in Extend) projecting_stable: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   169
     "projecting C h F (stable (extend_set h A)) (stable A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   170
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   171
apply (rule projecting_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   172
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   173
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   174
lemma (in Extend) projecting_increasing: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   175
     "projecting C h F (increasing (func o f)) (increasing func)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   176
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   177
apply (blast intro: project_increasing_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   178
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   179
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   180
lemma (in Extend) extending_UNIV: "extending C h F UNIV Y"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   181
apply (simp (no_asm) add: extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   182
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   183
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   184
lemma (in Extend) extending_constrains: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   185
     "extending (%G. UNIV) h F (extend_set h A co extend_set h B) (A co B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   186
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   187
apply (blast intro: project_constrains_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   188
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   189
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   190
lemma (in Extend) extending_stable: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   191
     "extending (%G. UNIV) h F (stable (extend_set h A)) (stable A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   192
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   193
apply (rule extending_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   194
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   195
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   196
lemma (in Extend) extending_increasing: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   197
     "extending (%G. UNIV) h F (increasing (func o f)) (increasing func)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   198
by (force simp only: extending_def Join_project_increasing)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   199
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   200
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   201
subsection{*Reachability and project*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   202
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   203
(*In practice, C = reachable(...): the inclusion is equality*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   204
lemma (in Extend) reachable_imp_reachable_project:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   205
     "[| reachable (extend h F\<squnion>G) \<subseteq> C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   206
         z \<in> reachable (extend h F\<squnion>G) |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   207
      ==> f z \<in> reachable (F\<squnion>project h C G)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   208
apply (erule reachable.induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   209
apply (force intro!: reachable.Init simp add: split_extended_all, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   210
 apply (rule_tac act = x in reachable.Acts)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   211
 apply auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   212
 apply (erule extend_act_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   213
apply (rule_tac act1 = "Restrict C act"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   214
       in project_act_I [THEN [3] reachable.Acts], auto) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   215
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   216
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   217
lemma (in Extend) project_Constrains_D: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   218
     "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Co B   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   219
      ==> extend h F\<squnion>G \<in> (extend_set h A) Co (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   220
apply (unfold Constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   221
apply (simp del: Join_constrains
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   222
            add: Join_project_constrains, clarify)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   223
apply (erule constrains_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   224
apply (auto intro: reachable_imp_reachable_project)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   225
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   226
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   227
lemma (in Extend) project_Stable_D: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   228
     "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Stable A   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   229
      ==> extend h F\<squnion>G \<in> Stable (extend_set h A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   230
apply (unfold Stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   231
apply (simp (no_asm_simp) add: project_Constrains_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   232
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   233
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   234
lemma (in Extend) project_Always_D: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   235
     "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Always A   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   236
      ==> extend h F\<squnion>G \<in> Always (extend_set h A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   237
apply (unfold Always_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   238
apply (force intro: reachable.Init simp add: project_Stable_D split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   239
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   240
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   241
lemma (in Extend) project_Increasing_D: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   242
     "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Increasing func   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   243
      ==> extend h F\<squnion>G \<in> Increasing (func o f)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   244
apply (unfold Increasing_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   245
apply (subst extend_set_eq_Collect [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   246
apply (simp (no_asm_simp) add: project_Stable_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   247
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   248
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   249
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   250
subsection{*Converse results for weak safety: benefits of the argument C *}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   251
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   252
(*In practice, C = reachable(...): the inclusion is equality*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   253
lemma (in Extend) reachable_project_imp_reachable:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   254
     "[| C \<subseteq> reachable(extend h F\<squnion>G);    
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   255
         x \<in> reachable (F\<squnion>project h C G) |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   256
      ==> \<exists>y. h(x,y) \<in> reachable (extend h F\<squnion>G)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   257
apply (erule reachable.induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   258
apply  (force intro: reachable.Init)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   259
apply (auto simp add: project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   260
apply (force del: Id_in_Acts intro: reachable.Acts extend_act_D)+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   261
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   262
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   263
lemma (in Extend) project_set_reachable_extend_eq:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   264
     "project_set h (reachable (extend h F\<squnion>G)) =  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   265
      reachable (F\<squnion>project h (reachable (extend h F\<squnion>G)) G)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   266
by (auto dest: subset_refl [THEN reachable_imp_reachable_project] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   267
               subset_refl [THEN reachable_project_imp_reachable])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   268
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   269
(*UNUSED*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   270
lemma (in Extend) reachable_extend_Join_subset:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   271
     "reachable (extend h F\<squnion>G) \<subseteq> C   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   272
      ==> reachable (extend h F\<squnion>G) \<subseteq>  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   273
          extend_set h (reachable (F\<squnion>project h C G))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   274
apply (auto dest: reachable_imp_reachable_project)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   275
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   276
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   277
lemma (in Extend) project_Constrains_I: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   278
     "extend h F\<squnion>G \<in> (extend_set h A) Co (extend_set h B)   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   279
      ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Co B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   280
apply (unfold Constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   281
apply (simp del: Join_constrains
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   282
            add: Join_project_constrains extend_set_Int_distrib)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   283
apply (rule conjI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   284
 prefer 2 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   285
 apply (force elim: constrains_weaken_L
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   286
              dest!: extend_constrains_project_set
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   287
                     subset_refl [THEN reachable_project_imp_reachable])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   288
apply (blast intro: constrains_weaken_L)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   289
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   290
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   291
lemma (in Extend) project_Stable_I: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   292
     "extend h F\<squnion>G \<in> Stable (extend_set h A)   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   293
      ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Stable A"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   294
apply (unfold Stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   295
apply (simp (no_asm_simp) add: project_Constrains_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   296
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   297
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   298
lemma (in Extend) project_Always_I: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   299
     "extend h F\<squnion>G \<in> Always (extend_set h A)   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   300
      ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Always A"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   301
apply (unfold Always_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   302
apply (auto simp add: project_Stable_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   303
apply (unfold extend_set_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   304
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   305
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   306
lemma (in Extend) project_Increasing_I: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   307
    "extend h F\<squnion>G \<in> Increasing (func o f)   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   308
     ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Increasing func"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   309
apply (unfold Increasing_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   310
apply (simp (no_asm_simp) add: extend_set_eq_Collect project_Stable_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   311
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   312
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   313
lemma (in Extend) project_Constrains:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   314
     "(F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Co B)  =   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   315
      (extend h F\<squnion>G \<in> (extend_set h A) Co (extend_set h B))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   316
apply (blast intro: project_Constrains_I project_Constrains_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   317
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   318
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   319
lemma (in Extend) project_Stable: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   320
     "(F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Stable A)  =   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   321
      (extend h F\<squnion>G \<in> Stable (extend_set h A))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   322
apply (unfold Stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   323
apply (rule project_Constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   324
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   325
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   326
lemma (in Extend) project_Increasing: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   327
   "(F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Increasing func)  =  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   328
    (extend h F\<squnion>G \<in> Increasing (func o f))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   329
apply (simp (no_asm_simp) add: Increasing_def project_Stable extend_set_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   330
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   331
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   332
subsection{*A lot of redundant theorems: all are proved to facilitate reasoning
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   333
    about guarantees.*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   334
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   335
lemma (in Extend) projecting_Constrains: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   336
     "projecting (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   337
                 (extend_set h A Co extend_set h B) (A Co B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   338
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   339
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   340
apply (blast intro: project_Constrains_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   341
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   342
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   343
lemma (in Extend) projecting_Stable: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   344
     "projecting (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   345
                 (Stable (extend_set h A)) (Stable A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   346
apply (unfold Stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   347
apply (rule projecting_Constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   348
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   349
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   350
lemma (in Extend) projecting_Always: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   351
     "projecting (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   352
                 (Always (extend_set h A)) (Always A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   353
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   354
apply (blast intro: project_Always_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   355
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   356
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   357
lemma (in Extend) projecting_Increasing: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   358
     "projecting (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   359
                 (Increasing (func o f)) (Increasing func)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   360
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   361
apply (blast intro: project_Increasing_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   362
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   363
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   364
lemma (in Extend) extending_Constrains: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   365
     "extending (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   366
                  (extend_set h A Co extend_set h B) (A Co B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   367
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   368
apply (blast intro: project_Constrains_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   369
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   370
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   371
lemma (in Extend) extending_Stable: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   372
     "extending (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   373
                  (Stable (extend_set h A)) (Stable A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   374
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   375
apply (blast intro: project_Stable_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   376
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   377
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   378
lemma (in Extend) extending_Always: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   379
     "extending (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   380
                  (Always (extend_set h A)) (Always A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   381
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   382
apply (blast intro: project_Always_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   383
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   384
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   385
lemma (in Extend) extending_Increasing: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   386
     "extending (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   387
                  (Increasing (func o f)) (Increasing func)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   388
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   389
apply (blast intro: project_Increasing_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   390
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   391
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   392
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   393
subsection{*leadsETo in the precondition (??)*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   394
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   395
subsubsection{*transient*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   396
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   397
lemma (in Extend) transient_extend_set_imp_project_transient: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   398
     "[| G \<in> transient (C \<inter> extend_set h A);  G \<in> stable C |]   
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   399
      ==> project h C G \<in> transient (project_set h C \<inter> A)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   400
apply (auto simp add: transient_def Domain_project_act)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   401
apply (subgoal_tac "act `` (C \<inter> extend_set h A) \<subseteq> - extend_set h A")
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   402
 prefer 2
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   403
 apply (simp add: stable_def constrains_def, blast) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   404
(*back to main goal*)
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   405
apply (erule_tac V = "?AA \<subseteq> -C \<union> ?BB" in thin_rl)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   406
apply (drule bspec, assumption) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   407
apply (simp add: extend_set_def project_act_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   408
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   409
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   410
(*converse might hold too?*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   411
lemma (in Extend) project_extend_transient_D: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   412
     "project h C (extend h F) \<in> transient (project_set h C \<inter> D)  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   413
      ==> F \<in> transient (project_set h C \<inter> D)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   414
apply (simp add: transient_def Domain_project_act, safe)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   415
apply blast+
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   416
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   417
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   418
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   419
subsubsection{*ensures -- a primitive combining progress with safety*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   420
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   421
(*Used to prove project_leadsETo_I*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   422
lemma (in Extend) ensures_extend_set_imp_project_ensures:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   423
     "[| extend h F \<in> stable C;  G \<in> stable C;   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   424
         extend h F\<squnion>G \<in> A ensures B;  A-B = C \<inter> extend_set h D |]   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   425
      ==> F\<squnion>project h C G   
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   426
            \<in> (project_set h C \<inter> project_set h A) ensures (project_set h B)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   427
apply (simp add: ensures_def project_constrains Join_transient extend_transient,
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   428
       clarify)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   429
apply (intro conjI) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   430
(*first subgoal*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   431
apply (blast intro: extend_stable_project_set 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   432
                  [THEN stableD, THEN constrains_Int, THEN constrains_weaken] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   433
             dest!: extend_constrains_project_set equalityD1)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   434
(*2nd subgoal*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   435
apply (erule stableD [THEN constrains_Int, THEN constrains_weaken])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   436
    apply assumption
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   437
   apply (simp (no_asm_use) add: extend_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   438
   apply blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   439
 apply (simp add: extend_set_Int_distrib extend_set_Un_distrib)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   440
 apply (blast intro!: extend_set_project_set [THEN subsetD], blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   441
(*The transient part*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   442
apply auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   443
 prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   444
 apply (force dest!: equalityD1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   445
              intro: transient_extend_set_imp_project_transient
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   446
                         [THEN transient_strengthen])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   447
apply (simp (no_asm_use) add: Int_Diff)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   448
apply (force dest!: equalityD1 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   449
             intro: transient_extend_set_imp_project_transient 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   450
               [THEN project_extend_transient_D, THEN transient_strengthen])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   451
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   452
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   453
text{*Transferring a transient property upwards*}
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   454
lemma (in Extend) project_transient_extend_set:
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   455
     "project h C G \<in> transient (project_set h C \<inter> A - B)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   456
      ==> G \<in> transient (C \<inter> extend_set h A - extend_set h B)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   457
apply (simp add: transient_def project_set_def extend_set_def project_act_def)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   458
apply (elim disjE bexE)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   459
 apply (rule_tac x=Id in bexI)  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   460
  apply (blast intro!: rev_bexI )+
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   461
done
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   462
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   463
lemma (in Extend) project_unless2 [rule_format]:
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   464
     "[| G \<in> stable C;  project h C G \<in> (project_set h C \<inter> A) unless B |]  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   465
      ==> G \<in> (C \<inter> extend_set h A) unless (extend_set h B)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   466
by (auto dest: stable_constrains_Int intro: constrains_weaken
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   467
         simp add: unless_def project_constrains Diff_eq Int_assoc 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   468
                   Int_extend_set_lemma)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   469
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   470
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   471
lemma (in Extend) extend_unless:
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   472
   "[|extend h F \<in> stable C; F \<in> A unless B|]
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   473
    ==> extend h F \<in> C \<inter> extend_set h A unless extend_set h B"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   474
apply (simp add: unless_def stable_def)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   475
apply (drule constrains_Int) 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   476
apply (erule extend_constrains [THEN iffD2]) 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   477
apply (erule constrains_weaken, blast) 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   478
apply blast 
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   479
done
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   480
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   481
(*Used to prove project_leadsETo_D*)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   482
lemma (in Extend) Join_project_ensures [rule_format]:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   483
     "[| extend h F\<squnion>G \<in> stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   484
         F\<squnion>project h C G \<in> A ensures B |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   485
      ==> extend h F\<squnion>G \<in> (C \<inter> extend_set h A) ensures (extend_set h B)"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   486
apply (auto simp add: ensures_eq extend_unless project_unless)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   487
apply (blast intro:  extend_transient [THEN iffD2] transient_strengthen)  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   488
apply (blast intro: project_transient_extend_set transient_strengthen)  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   489
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   490
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   491
text{*Lemma useful for both STRONG and WEAK progress, but the transient
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   492
    condition's very strong*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   493
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   494
(*The strange induction formula allows induction over the leadsTo
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   495
  assumption's non-atomic precondition*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   496
lemma (in Extend) PLD_lemma:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   497
     "[| extend h F\<squnion>G \<in> stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   498
         F\<squnion>project h C G \<in> (project_set h C \<inter> A) leadsTo B |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   499
      ==> extend h F\<squnion>G \<in>  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   500
          C \<inter> extend_set h (project_set h C \<inter> A) leadsTo (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   501
apply (erule leadsTo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   502
  apply (blast intro: leadsTo_Basis Join_project_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   503
 apply (blast intro: psp_stable2 [THEN leadsTo_weaken_L] leadsTo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   504
apply (simp del: UN_simps add: Int_UN_distrib leadsTo_UN extend_set_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   505
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   506
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   507
lemma (in Extend) project_leadsTo_D_lemma:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   508
     "[| extend h F\<squnion>G \<in> stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   509
         F\<squnion>project h C G \<in> (project_set h C \<inter> A) leadsTo B |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   510
      ==> extend h F\<squnion>G \<in> (C \<inter> extend_set h A) leadsTo (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   511
apply (rule PLD_lemma [THEN leadsTo_weaken])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   512
apply (auto simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   513
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   514
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   515
lemma (in Extend) Join_project_LeadsTo:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   516
     "[| C = (reachable (extend h F\<squnion>G));  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   517
         F\<squnion>project h C G \<in> A LeadsTo B |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   518
      ==> extend h F\<squnion>G \<in> (extend_set h A) LeadsTo (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   519
by (simp del: Join_stable    add: LeadsTo_def project_leadsTo_D_lemma
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   520
                                  project_set_reachable_extend_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   521
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   522
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   523
subsection{*Towards the theorem @{text project_Ensures_D}*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   524
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   525
lemma (in Extend) project_ensures_D_lemma:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   526
     "[| G \<in> stable ((C \<inter> extend_set h A) - (extend_set h B));   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   527
         F\<squnion>project h C G \<in> (project_set h C \<inter> A) ensures B;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   528
         extend h F\<squnion>G \<in> stable C |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   529
      ==> extend h F\<squnion>G \<in> (C \<inter> extend_set h A) ensures (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   530
(*unless*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   531
apply (auto intro!: project_unless2 [unfolded unless_def] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   532
            intro: project_extend_constrains_I 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   533
            simp add: ensures_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   534
(*transient*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   535
(*A G-action cannot occur*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   536
 prefer 2
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   537
 apply (blast intro: project_transient_extend_set) 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   538
(*An F-action*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   539
apply (force elim!: extend_transient [THEN iffD2, THEN transient_strengthen]
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   540
             simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   541
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   542
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   543
lemma (in Extend) project_ensures_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   544
     "[| F\<squnion>project h UNIV G \<in> A ensures B;   
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   545
         G \<in> stable (extend_set h A - extend_set h B) |]  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   546
      ==> extend h F\<squnion>G \<in> (extend_set h A) ensures (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   547
apply (rule project_ensures_D_lemma [of _ UNIV, THEN revcut_rl], auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   548
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   549
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   550
lemma (in Extend) project_Ensures_D: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   551
     "[| F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Ensures B;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   552
         G \<in> stable (reachable (extend h F\<squnion>G) \<inter> extend_set h A -  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   553
                     extend_set h B) |]  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   554
      ==> extend h F\<squnion>G \<in> (extend_set h A) Ensures (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   555
apply (unfold Ensures_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   556
apply (rule project_ensures_D_lemma [THEN revcut_rl])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   557
apply (auto simp add: project_set_reachable_extend_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   558
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   559
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   560
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   561
subsection{*Guarantees*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   562
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   563
lemma (in Extend) project_act_Restrict_subset_project_act:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   564
     "project_act h (Restrict C act) \<subseteq> project_act h act"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   565
apply (auto simp add: project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   566
done
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
   567
                                           
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
   568
                                                           
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   569
lemma (in Extend) subset_closed_ok_extend_imp_ok_project:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   570
     "[| extend h F ok G; subset_closed (AllowedActs F) |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   571
      ==> F ok project h C G"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   572
apply (auto simp add: ok_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   573
apply (rename_tac act) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   574
apply (drule subsetD, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   575
apply (rule_tac x = "Restrict C  (extend_act h act)" in rev_image_eqI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   576
apply simp +
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   577
apply (cut_tac project_act_Restrict_subset_project_act)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   578
apply (auto simp add: subset_closed_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   579
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   580
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   581
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   582
(*Weak precondition and postcondition
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   583
  Not clear that it has a converse [or that we want one!]*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   584
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   585
(*The raw version; 3rd premise could be weakened by adding the
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   586
  precondition extend h F\<squnion>G \<in> X' *)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   587
lemma (in Extend) project_guarantees_raw:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   588
 assumes xguary:  "F \<in> X guarantees Y"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   589
     and closed:  "subset_closed (AllowedActs F)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   590
     and project: "!!G. extend h F\<squnion>G \<in> X' 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   591
                        ==> F\<squnion>project h (C G) G \<in> X"
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   592
     and extend:  "!!G. [| F\<squnion>project h (C G) G \<in> Y |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   593
                        ==> extend h F\<squnion>G \<in> Y'"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   594
 shows "extend h F \<in> X' guarantees Y'"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   595
apply (rule xguary [THEN guaranteesD, THEN extend, THEN guaranteesI])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   596
apply (blast intro: closed subset_closed_ok_extend_imp_ok_project)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   597
apply (erule project)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   598
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   599
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   600
lemma (in Extend) project_guarantees:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   601
     "[| F \<in> X guarantees Y;  subset_closed (AllowedActs F);  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   602
         projecting C h F X' X;  extending C h F Y' Y |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   603
      ==> extend h F \<in> X' guarantees Y'"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   604
apply (rule guaranteesI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   605
apply (auto simp add: guaranteesD projecting_def extending_def
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   606
                      subset_closed_ok_extend_imp_ok_project)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   607
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   608
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   609
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   610
(*It seems that neither "guarantees" law can be proved from the other.*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   611
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   612
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   613
subsection{*guarantees corollaries*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   614
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   615
subsubsection{*Some could be deleted: the required versions are easy to prove*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   616
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   617
lemma (in Extend) extend_guar_increasing:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   618
     "[| F \<in> UNIV guarantees increasing func;   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   619
         subset_closed (AllowedActs F) |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   620
      ==> extend h F \<in> X' guarantees increasing (func o f)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   621
apply (erule project_guarantees)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   622
apply (rule_tac [3] extending_increasing)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   623
apply (rule_tac [2] projecting_UNIV, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   624
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   625
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   626
lemma (in Extend) extend_guar_Increasing:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   627
     "[| F \<in> UNIV guarantees Increasing func;   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   628
         subset_closed (AllowedActs F) |]  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   629
      ==> extend h F \<in> X' guarantees Increasing (func o f)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   630
apply (erule project_guarantees)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   631
apply (rule_tac [3] extending_Increasing)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   632
apply (rule_tac [2] projecting_UNIV, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   633
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   634
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   635
lemma (in Extend) extend_guar_Always:
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   636
     "[| F \<in> Always A guarantees Always B;   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   637
         subset_closed (AllowedActs F) |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   638
      ==> extend h F                    
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   639
            \<in> Always(extend_set h A) guarantees Always(extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   640
apply (erule project_guarantees)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   641
apply (rule_tac [3] extending_Always)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   642
apply (rule_tac [2] projecting_Always, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   643
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   644
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   645
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   646
subsubsection{*Guarantees with a leadsTo postcondition*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   647
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   648
lemma (in Extend) project_leadsTo_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   649
     "F\<squnion>project h UNIV G \<in> A leadsTo B
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   650
      ==> extend h F\<squnion>G \<in> (extend_set h A) leadsTo (extend_set h B)"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   651
apply (rule_tac C1 = UNIV in project_leadsTo_D_lemma [THEN leadsTo_weaken], auto)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   652
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   653
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   654
lemma (in Extend) project_LeadsTo_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   655
     "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A LeadsTo B   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   656
       ==> extend h F\<squnion>G \<in> (extend_set h A) LeadsTo (extend_set h B)"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   657
apply (rule refl [THEN Join_project_LeadsTo], auto)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   658
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   659
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   660
lemma (in Extend) extending_leadsTo: 
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   661
     "extending (%G. UNIV) h F  
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   662
                (extend_set h A leadsTo extend_set h B) (A leadsTo B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   663
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   664
apply (blast intro: project_leadsTo_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   665
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   666
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   667
lemma (in Extend) extending_LeadsTo: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   668
     "extending (%G. reachable (extend h F\<squnion>G)) h F  
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   669
                (extend_set h A LeadsTo extend_set h B) (A LeadsTo B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   670
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   671
apply (blast intro: project_LeadsTo_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   672
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10064
diff changeset
   673
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   674
end