src/HOL/UNITY/ELT.thy
author steckerm
Sat, 20 Sep 2014 10:44:24 +0200
changeset 58407 111d801b5d5d
parent 46912 e0cd5c4df8e6
child 58889 5b7a9633cfa8
permissions -rw-r--r--
Changed proof method to auto for custom Waldmeister lemma
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26806
diff changeset
     1
(*  Title:      HOL/UNITY/ELT.thy
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     3
    Copyright   1999  University of Cambridge
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     4
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     5
leadsTo strengthened with a specification of the allowable sets transient parts
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
     6
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
     7
TRY INSTEAD (to get rid of the {} and to gain strong induction)
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
     8
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
     9
  elt :: "['a set set, 'a program, 'a set] => ('a set) set"
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    10
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    11
inductive "elt CC F B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    12
  intros 
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    13
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    14
    Weaken:  "A <= B ==> A : elt CC F B"
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    15
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    16
    ETrans:  "[| F : A ensures A';  A-A' : CC;  A' : elt CC F B |]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26806
diff changeset
    17
              ==> A : elt CC F B"
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    18
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    19
    Union:  "{A. A: S} : Pow (elt CC F B) ==> (Union S) : elt CC F B"
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    20
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8072
diff changeset
    21
  monos Pow_mono
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    22
*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    23
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    24
header{*Progress Under Allowable Sets*}
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    25
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13836
diff changeset
    26
theory ELT imports Project begin
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    27
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    28
inductive_set
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    29
  (*LEADS-TO constant for the inductive definition*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    30
  elt :: "['a set set, 'a program] => ('a set * 'a set) set"
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    31
  for CC :: "'a set set" and F :: "'a program"
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    32
 where
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    33
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    34
   Basis:  "[| F : A ensures B;  A-B : (insert {} CC) |] ==> (A,B) : elt CC F"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    35
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    36
 | Trans:  "[| (A,B) : elt CC F;  (B,C) : elt CC F |] ==> (A,C) : elt CC F"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    37
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    38
 | Union:  "ALL A: S. (A,B) : elt CC F ==> (Union S, B) : elt CC F"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    39
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    40
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    41
definition  
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
    42
  (*the set of all sets determined by f alone*)
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
    43
  givenBy :: "['a => 'b] => 'a set set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    44
  where "givenBy f = range (%B. f-` B)"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    45
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    46
definition
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    47
  (*visible version of the LEADS-TO relation*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    48
  leadsETo :: "['a set, 'a set set, 'a set] => 'a program set"
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    49
                                        ("(3_/ leadsTo[_]/ _)" [80,0,80] 80)
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    50
  where "leadsETo A CC B = {F. (A,B) : elt CC F}"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    51
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    52
definition
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    53
  LeadsETo :: "['a set, 'a set set, 'a set] => 'a program set"
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    54
                                        ("(3_/ LeadsTo[_]/ _)" [80,0,80] 80)
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    55
  where "LeadsETo A CC B =
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10293
diff changeset
    56
      {F. F : (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] B}"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    57
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    58
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    59
(*** givenBy ***)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    60
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    61
lemma givenBy_id [simp]: "givenBy id = UNIV"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    62
by (unfold givenBy_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    63
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    64
lemma givenBy_eq_all: "(givenBy v) = {A. ALL x:A. ALL y. v x = v y --> y: A}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    65
apply (unfold givenBy_def, safe)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    66
apply (rule_tac [2] x = "v ` ?u" in image_eqI, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    67
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    68
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    69
lemma givenByI: "(!!x y. [| x:A;  v x = v y |] ==> y: A) ==> A: givenBy v"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    70
by (subst givenBy_eq_all, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    71
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    72
lemma givenByD: "[| A: givenBy v;  x:A;  v x = v y |] ==> y: A"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    73
by (unfold givenBy_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    74
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    75
lemma empty_mem_givenBy [iff]: "{} : givenBy v"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    76
by (blast intro!: givenByI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    77
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    78
lemma givenBy_imp_eq_Collect: "A: givenBy v ==> EX P. A = {s. P(v s)}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    79
apply (rule_tac x = "%n. EX s. v s = n & s : A" in exI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    80
apply (simp (no_asm_use) add: givenBy_eq_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    81
apply blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    82
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    83
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    84
lemma Collect_mem_givenBy: "{s. P(v s)} : givenBy v"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    85
by (unfold givenBy_def, best)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    86
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    87
lemma givenBy_eq_Collect: "givenBy v = {A. EX P. A = {s. P(v s)}}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    88
by (blast intro: Collect_mem_givenBy givenBy_imp_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    89
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    90
(*preserving v preserves properties given by v*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    91
lemma preserves_givenBy_imp_stable:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    92
     "[| F : preserves v;  D : givenBy v |] ==> F : stable D"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    93
by (force simp add: preserves_subset_stable [THEN subsetD] givenBy_eq_Collect)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    94
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    95
lemma givenBy_o_subset: "givenBy (w o v) <= givenBy v"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    96
apply (simp (no_asm) add: givenBy_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    97
apply best 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    98
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    99
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   100
lemma givenBy_DiffI:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   101
     "[| A : givenBy v;  B : givenBy v |] ==> A-B : givenBy v"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   102
apply (simp (no_asm_use) add: givenBy_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   103
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   104
apply (rule_tac x = "%z. ?R z & ~ ?Q z" in exI)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26349
diff changeset
   105
unfolding set_diff_eq
26349
7f5a2f6d9119 tuned proof
haftmann
parents: 23767
diff changeset
   106
apply auto
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   107
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   108
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   109
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   110
(** Standard leadsTo rules **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   111
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   112
lemma leadsETo_Basis [intro]: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   113
     "[| F: A ensures B;  A-B: insert {} CC |] ==> F : A leadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   114
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   115
apply (blast intro: elt.Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   116
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   117
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   118
lemma leadsETo_Trans: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   119
     "[| F : A leadsTo[CC] B;  F : B leadsTo[CC] C |] ==> F : A leadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   120
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   121
apply (blast intro: elt.Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   122
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   123
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   124
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   125
(*Useful with cancellation, disjunction*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   126
lemma leadsETo_Un_duplicate:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   127
     "F : A leadsTo[CC] (A' Un A') ==> F : A leadsTo[CC] A'"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   128
by (simp add: Un_ac)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   129
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   130
lemma leadsETo_Un_duplicate2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   131
     "F : A leadsTo[CC] (A' Un C Un C) ==> F : A leadsTo[CC] (A' Un C)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   132
by (simp add: Un_ac)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   133
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   134
(*The Union introduction rule as we should have liked to state it*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   135
lemma leadsETo_Union:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   136
    "(!!A. A : S ==> F : A leadsTo[CC] B) ==> F : (Union S) leadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   137
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   138
apply (blast intro: elt.Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   139
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   140
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   141
lemma leadsETo_UN:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   142
    "(!!i. i : I ==> F : (A i) leadsTo[CC] B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   143
     ==> F : (UN i:I. A i) leadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   144
apply (subst Union_image_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   145
apply (blast intro: leadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   146
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   147
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   148
(*The INDUCTION rule as we should have liked to state it*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   149
lemma leadsETo_induct:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   150
  "[| F : za leadsTo[CC] zb;   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   151
      !!A B. [| F : A ensures B;  A-B : insert {} CC |] ==> P A B;  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   152
      !!A B C. [| F : A leadsTo[CC] B; P A B; F : B leadsTo[CC] C; P B C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   153
               ==> P A C;  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   154
      !!B S. ALL A:S. F : A leadsTo[CC] B & P A B ==> P (Union S) B  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   155
   |] ==> P za zb"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   156
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   157
apply (drule CollectD) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   158
apply (erule elt.induct, blast+)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   159
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   160
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   161
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   162
(** New facts involving leadsETo **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   163
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   164
lemma leadsETo_mono: "CC' <= CC ==> (A leadsTo[CC'] B) <= (A leadsTo[CC] B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   165
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   166
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   167
prefer 3 apply (blast intro: leadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   168
prefer 2 apply (blast intro: leadsETo_Trans)
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45605
diff changeset
   169
apply blast
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   170
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   171
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   172
lemma leadsETo_Trans_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   173
     "[| F : A leadsTo[CC] B;  F : B leadsTo[DD] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   174
      ==> F : A leadsTo[CC Un DD] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   175
by (blast intro: leadsETo_mono [THEN subsetD] leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   176
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   177
lemma leadsETo_Union_Int:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   178
 "(!!A. A : S ==> F : (A Int C) leadsTo[CC] B) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   179
  ==> F : (Union S Int C) leadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   180
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   181
apply (simp only: Int_Union_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   182
apply (blast intro: elt.Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   183
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   184
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   185
(*Binary union introduction rule*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   186
lemma leadsETo_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   187
     "[| F : A leadsTo[CC] C; F : B leadsTo[CC] C |] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   188
      ==> F : (A Un B) leadsTo[CC] C"
44106
0e018cbcc0de tuned proofs
haftmann
parents: 36866
diff changeset
   189
  using leadsETo_Union [of "{A, B}" F CC C] by auto
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   190
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   191
lemma single_leadsETo_I:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   192
     "(!!x. x : A ==> F : {x} leadsTo[CC] B) ==> F : A leadsTo[CC] B"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   193
by (subst UN_singleton [symmetric], rule leadsETo_UN, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   194
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   195
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   196
lemma subset_imp_leadsETo: "A<=B ==> F : A leadsTo[CC] B"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   197
by (simp add: subset_imp_ensures [THEN leadsETo_Basis] 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   198
              Diff_eq_empty_iff [THEN iffD2])
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   199
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   200
lemmas empty_leadsETo = empty_subsetI [THEN subset_imp_leadsETo, simp]
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   201
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   202
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   203
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   204
(** Weakening laws **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   205
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   206
lemma leadsETo_weaken_R:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   207
     "[| F : A leadsTo[CC] A';  A'<=B' |] ==> F : A leadsTo[CC] B'"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   208
by (blast intro: subset_imp_leadsETo leadsETo_Trans)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   209
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   210
lemma leadsETo_weaken_L:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   211
     "[| F : A leadsTo[CC] A'; B<=A |] ==> F : B leadsTo[CC] A'"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   212
by (blast intro: leadsETo_Trans subset_imp_leadsETo)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   213
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   214
(*Distributes over binary unions*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   215
lemma leadsETo_Un_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   216
     "F : (A Un B) leadsTo[CC] C  =   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   217
      (F : A leadsTo[CC] C & F : B leadsTo[CC] C)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   218
by (blast intro: leadsETo_Un leadsETo_weaken_L)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   219
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   220
lemma leadsETo_UN_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   221
     "F : (UN i:I. A i) leadsTo[CC] B  =   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   222
      (ALL i : I. F : (A i) leadsTo[CC] B)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   223
by (blast intro: leadsETo_UN leadsETo_weaken_L)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   224
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   225
lemma leadsETo_Union_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   226
     "F : (Union S) leadsTo[CC] B  =  (ALL A : S. F : A leadsTo[CC] B)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   227
by (blast intro: leadsETo_Union leadsETo_weaken_L)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   228
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   229
lemma leadsETo_weaken:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   230
     "[| F : A leadsTo[CC'] A'; B<=A; A'<=B';  CC' <= CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   231
      ==> F : B leadsTo[CC] B'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   232
apply (drule leadsETo_mono [THEN subsetD], assumption)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   233
apply (blast del: subsetCE 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   234
             intro: leadsETo_weaken_R leadsETo_weaken_L leadsETo_Trans)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   235
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   236
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   237
lemma leadsETo_givenBy:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   238
     "[| F : A leadsTo[CC] A';  CC <= givenBy v |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   239
      ==> F : A leadsTo[givenBy v] A'"
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45605
diff changeset
   240
by (blast intro: leadsETo_weaken)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   241
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   242
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   243
(*Set difference*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   244
lemma leadsETo_Diff:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   245
     "[| F : (A-B) leadsTo[CC] C; F : B leadsTo[CC] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   246
      ==> F : A leadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   247
by (blast intro: leadsETo_Un leadsETo_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   248
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   249
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   250
(*Binary union version*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   251
lemma leadsETo_Un_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   252
     "[| F : A leadsTo[CC] A';  F : B leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   253
      ==> F : (A Un B) leadsTo[CC] (A' Un B')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   254
by (blast intro: leadsETo_Un leadsETo_weaken_R)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   255
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   256
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   257
(** The cancellation law **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   258
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   259
lemma leadsETo_cancel2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   260
     "[| F : A leadsTo[CC] (A' Un B); F : B leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   261
      ==> F : A leadsTo[CC] (A' Un B')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   262
by (blast intro: leadsETo_Un_Un subset_imp_leadsETo leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   263
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   264
lemma leadsETo_cancel1:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   265
     "[| F : A leadsTo[CC] (B Un A'); F : B leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   266
    ==> F : A leadsTo[CC] (B' Un A')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   267
apply (simp add: Un_commute)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   268
apply (blast intro!: leadsETo_cancel2)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   269
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   270
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   271
lemma leadsETo_cancel_Diff1:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   272
     "[| F : A leadsTo[CC] (B Un A'); F : (B-A') leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   273
    ==> F : A leadsTo[CC] (B' Un A')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   274
apply (rule leadsETo_cancel1)
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   275
 prefer 2 apply assumption
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   276
apply simp_all
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   277
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   278
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   279
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   280
(** PSP: Progress-Safety-Progress **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   281
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   282
(*Special case of PSP: Misra's "stable conjunction"*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   283
lemma e_psp_stable: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   284
   "[| F : A leadsTo[CC] A';  F : stable B;  ALL C:CC. C Int B : CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   285
    ==> F : (A Int B) leadsTo[CC] (A' Int B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   286
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   287
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   288
prefer 3 apply (blast intro: leadsETo_Union_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   289
prefer 2 apply (blast intro: leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   290
apply (rule leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   291
prefer 2 apply (force simp add: Diff_Int_distrib2 [symmetric])
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   292
apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   293
                 Int_Un_distrib2 [symmetric])
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   294
apply (blast intro: transient_strengthen constrains_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   295
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   296
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   297
lemma e_psp_stable2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   298
     "[| F : A leadsTo[CC] A'; F : stable B;  ALL C:CC. C Int B : CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   299
      ==> F : (B Int A) leadsTo[CC] (B Int A')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   300
by (simp (no_asm_simp) add: e_psp_stable Int_ac)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   301
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   302
lemma e_psp:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   303
     "[| F : A leadsTo[CC] A'; F : B co B';   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   304
         ALL C:CC. C Int B Int B' : CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   305
      ==> F : (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   306
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   307
prefer 3 apply (blast intro: leadsETo_Union_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   308
(*Transitivity case has a delicate argument involving "cancellation"*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   309
apply (rule_tac [2] leadsETo_Un_duplicate2)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   310
apply (erule_tac [2] leadsETo_cancel_Diff1)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   311
prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   312
 apply (simp add: Int_Diff Diff_triv)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   313
 apply (blast intro: leadsETo_weaken_L dest: constrains_imp_subset)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   314
(*Basis case*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   315
apply (rule leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   316
apply (blast intro: psp_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   317
apply (subgoal_tac "A Int B' - (Ba Int B Un (B' - B)) = (A - Ba) Int B Int B'")
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   318
apply auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   319
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   320
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   321
lemma e_psp2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   322
     "[| F : A leadsTo[CC] A'; F : B co B';   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   323
         ALL C:CC. C Int B Int B' : CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   324
      ==> F : (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   325
by (simp add: e_psp Int_ac)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   326
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   327
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   328
(*** Special properties involving the parameter [CC] ***)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   329
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   330
(*??IS THIS NEEDED?? or is it just an example of what's provable??*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   331
lemma gen_leadsETo_imp_Join_leadsETo:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   332
     "[| F: (A leadsTo[givenBy v] B);  G : preserves v;   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   333
         F\<squnion>G : stable C |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   334
      ==> F\<squnion>G : ((C Int A) leadsTo[(%D. C Int D) ` givenBy v] B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   335
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   336
  prefer 3
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   337
  apply (subst Int_Union) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   338
  apply (blast intro: leadsETo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   339
prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   340
 apply (blast intro: e_psp_stable2 [THEN leadsETo_weaken_L] leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   341
apply (rule leadsETo_Basis)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   342
apply (auto simp add: Diff_eq_empty_iff [THEN iffD2] 
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45605
diff changeset
   343
                      Int_Diff ensures_def givenBy_eq_Collect)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   344
prefer 3 apply (blast intro: transient_strengthen)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   345
apply (drule_tac [2] P1 = P in preserves_subset_stable [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   346
apply (drule_tac P1 = P in preserves_subset_stable [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   347
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   348
apply (blast intro: constrains_Int [THEN constrains_weaken])+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   349
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   350
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   351
(**** Relationship with traditional "leadsTo", strong & weak ****)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   352
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   353
(** strong **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   354
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   355
lemma leadsETo_subset_leadsTo: "(A leadsTo[CC] B) <= (A leadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   356
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   357
apply (erule leadsETo_induct)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   358
  prefer 3 apply (blast intro: leadsTo_Union)
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   359
 prefer 2 apply (blast intro: leadsTo_Trans, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   360
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   361
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   362
lemma leadsETo_UNIV_eq_leadsTo: "(A leadsTo[UNIV] B) = (A leadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   363
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   364
apply (erule leadsETo_subset_leadsTo [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   365
(*right-to-left case*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   366
apply (erule leadsTo_induct)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   367
  prefer 3 apply (blast intro: leadsETo_Union)
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   368
 prefer 2 apply (blast intro: leadsETo_Trans, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   369
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   370
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   371
(**** weak ****)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   372
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   373
lemma LeadsETo_eq_leadsETo: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   374
     "A LeadsTo[CC] B =  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   375
        {F. F : (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   376
        (reachable F Int B)}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   377
apply (unfold LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   378
apply (blast dest: e_psp_stable2 intro: leadsETo_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   379
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   380
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   381
(*** Introduction rules: Basis, Trans, Union ***)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   382
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   383
lemma LeadsETo_Trans:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   384
     "[| F : A LeadsTo[CC] B;  F : B LeadsTo[CC] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   385
      ==> F : A LeadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   386
apply (simp add: LeadsETo_eq_leadsETo)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   387
apply (blast intro: leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   388
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   389
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   390
lemma LeadsETo_Union:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   391
     "(!!A. A : S ==> F : A LeadsTo[CC] B) ==> F : (Union S) LeadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   392
apply (simp add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   393
apply (subst Int_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   394
apply (blast intro: leadsETo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   395
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   396
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   397
lemma LeadsETo_UN:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   398
     "(!!i. i : I ==> F : (A i) LeadsTo[CC] B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   399
      ==> F : (UN i:I. A i) LeadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   400
apply (simp only: Union_image_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   401
apply (blast intro: LeadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   402
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   403
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   404
(*Binary union introduction rule*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   405
lemma LeadsETo_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   406
     "[| F : A LeadsTo[CC] C; F : B LeadsTo[CC] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   407
      ==> F : (A Un B) LeadsTo[CC] C"
44106
0e018cbcc0de tuned proofs
haftmann
parents: 36866
diff changeset
   408
  using LeadsETo_Union [of "{A, B}" F CC C] by auto
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   409
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   410
(*Lets us look at the starting state*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   411
lemma single_LeadsETo_I:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   412
     "(!!s. s : A ==> F : {s} LeadsTo[CC] B) ==> F : A LeadsTo[CC] B"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   413
by (subst UN_singleton [symmetric], rule LeadsETo_UN, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   414
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   415
lemma subset_imp_LeadsETo:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   416
     "A <= B ==> F : A LeadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   417
apply (simp (no_asm) add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   418
apply (blast intro: subset_imp_leadsETo)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   419
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   420
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44106
diff changeset
   421
lemmas empty_LeadsETo = empty_subsetI [THEN subset_imp_LeadsETo]
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   422
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   423
lemma LeadsETo_weaken_R:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   424
     "[| F : A LeadsTo[CC] A';  A' <= B' |] ==> F : A LeadsTo[CC] B'"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   425
apply (simp add: LeadsETo_def)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   426
apply (blast intro: leadsETo_weaken_R)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   427
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   428
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   429
lemma LeadsETo_weaken_L:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   430
     "[| F : A LeadsTo[CC] A';  B <= A |] ==> F : B LeadsTo[CC] A'"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   431
apply (simp add: LeadsETo_def)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   432
apply (blast intro: leadsETo_weaken_L)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   433
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   434
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   435
lemma LeadsETo_weaken:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   436
     "[| F : A LeadsTo[CC'] A';    
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   437
         B <= A;  A' <= B';  CC' <= CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   438
      ==> F : B LeadsTo[CC] B'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   439
apply (simp (no_asm_use) add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   440
apply (blast intro: leadsETo_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   441
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   442
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   443
lemma LeadsETo_subset_LeadsTo: "(A LeadsTo[CC] B) <= (A LeadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   444
apply (unfold LeadsETo_def LeadsTo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   445
apply (blast intro: leadsETo_subset_leadsTo [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   446
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   447
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   448
(*Postcondition can be strengthened to (reachable F Int B) *)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   449
lemma reachable_ensures:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   450
     "F : A ensures B ==> F : (reachable F Int A) ensures B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   451
apply (rule stable_ensures_Int [THEN ensures_weaken_R], auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   452
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   453
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   454
lemma lel_lemma:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   455
     "F : A leadsTo B ==> F : (reachable F Int A) leadsTo[Pow(reachable F)] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   456
apply (erule leadsTo_induct)
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45605
diff changeset
   457
  apply (blast intro: reachable_ensures)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   458
 apply (blast dest: e_psp_stable2 intro: leadsETo_Trans leadsETo_weaken_L)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   459
apply (subst Int_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   460
apply (blast intro: leadsETo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   461
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   462
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   463
lemma LeadsETo_UNIV_eq_LeadsTo: "(A LeadsTo[UNIV] B) = (A LeadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   464
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   465
apply (erule LeadsETo_subset_LeadsTo [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   466
(*right-to-left case*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   467
apply (unfold LeadsETo_def LeadsTo_def)
13836
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   468
apply (blast intro: lel_lemma [THEN leadsETo_weaken])
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   469
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   470
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   471
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   472
(**** EXTEND/PROJECT PROPERTIES ****)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   473
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   474
context Extend
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   475
begin
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   476
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   477
lemma givenBy_o_eq_extend_set:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   478
     "givenBy (v o f) = extend_set h ` (givenBy v)"
13836
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   479
apply (simp add: givenBy_eq_Collect)
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   480
apply (rule equalityI, best)
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   481
apply blast 
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   482
done
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   483
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   484
lemma givenBy_eq_extend_set: "givenBy f = range (extend_set h)"
13836
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   485
by (simp add: givenBy_eq_Collect, best)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   486
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   487
lemma extend_set_givenBy_I:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   488
     "D : givenBy v ==> extend_set h D : givenBy (v o f)"
13836
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   489
apply (simp (no_asm_use) add: givenBy_eq_all, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   490
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   491
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   492
lemma leadsETo_imp_extend_leadsETo:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   493
     "F : A leadsTo[CC] B  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   494
      ==> extend h F : (extend_set h A) leadsTo[extend_set h ` CC]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   495
                       (extend_set h B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   496
apply (erule leadsETo_induct)
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45605
diff changeset
   497
  apply (force intro: subset_imp_ensures 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   498
               simp add: extend_ensures extend_set_Diff_distrib [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   499
 apply (blast intro: leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   500
apply (simp add: leadsETo_UN extend_set_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   501
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   502
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   503
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   504
(*This version's stronger in the "ensures" precondition
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   505
  BUT there's no ensures_weaken_L*)
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   506
lemma Join_project_ensures_strong:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   507
     "[| project h C G ~: transient (project_set h C Int (A-B)) |  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   508
           project_set h C Int (A - B) = {};   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   509
         extend h F\<squnion>G : stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   510
         F\<squnion>project h C G : (project_set h C Int A) ensures B |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   511
      ==> extend h F\<squnion>G : (C Int extend_set h A) ensures (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   512
apply (subst Int_extend_set_lemma [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   513
apply (rule Join_project_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   514
apply (auto simp add: Int_Diff)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   515
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   516
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   517
(*NOT WORKING.  MODIFY AS IN Project.thy
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   518
lemma pld_lemma:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   519
     "[| extend h F\<squnion>G : stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   520
         F\<squnion>project h C G : (project_set h C Int A) leadsTo[(%D. project_set h C Int D)`givenBy v] B;   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   521
         G : preserves (v o f) |]  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   522
      ==> extend h F\<squnion>G :  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   523
            (C Int extend_set h (project_set h C Int A))  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   524
            leadsTo[(%D. C Int extend_set h D)`givenBy v]  (extend_set h B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   525
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   526
  prefer 3
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   527
  apply (simp del: UN_simps add: Int_UN_distrib leadsETo_UN extend_set_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   528
 prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   529
 apply (blast intro: e_psp_stable2 [THEN leadsETo_weaken_L] leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   530
txt{*Base case is hard*}
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   531
apply auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   532
apply (force intro: leadsETo_Basis subset_imp_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   533
apply (rule leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   534
 prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   535
 apply (simp add: Int_Diff Int_extend_set_lemma extend_set_Diff_distrib [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   536
apply (rule Join_project_ensures_strong)
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   537
apply (auto intro: project_stable_project_set simp add: Int_left_absorb)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   538
apply (simp (no_asm_simp) add: stable_ensures_Int [THEN ensures_weaken_R] Int_lower2 project_stable_project_set extend_stable_project_set)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   539
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   540
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   541
lemma project_leadsETo_D_lemma:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   542
     "[| extend h F\<squnion>G : stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   543
         F\<squnion>project h C G :  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   544
             (project_set h C Int A)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   545
             leadsTo[(%D. project_set h C Int D)`givenBy v] B;   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   546
         G : preserves (v o f) |]  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   547
      ==> extend h F\<squnion>G : (C Int extend_set h A)  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   548
            leadsTo[(%D. C Int extend_set h D)`givenBy v] (extend_set h B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   549
apply (rule pld_lemma [THEN leadsETo_weaken])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   550
apply (auto simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   551
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   552
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   553
lemma project_leadsETo_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   554
     "[| F\<squnion>project h UNIV G : A leadsTo[givenBy v] B;   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   555
         G : preserves (v o f) |]   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   556
      ==> extend h F\<squnion>G : (extend_set h A)  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   557
            leadsTo[givenBy (v o f)] (extend_set h B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   558
apply (cut_tac project_leadsETo_D_lemma [of _ _ UNIV], auto) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   559
apply (erule leadsETo_givenBy)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   560
apply (rule givenBy_o_eq_extend_set [THEN equalityD2])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   561
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   562
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   563
lemma project_LeadsETo_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   564
     "[| F\<squnion>project h (reachable (extend h F\<squnion>G)) G  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   565
             : A LeadsTo[givenBy v] B;   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   566
         G : preserves (v o f) |]  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   567
      ==> extend h F\<squnion>G :  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   568
            (extend_set h A) LeadsTo[givenBy (v o f)] (extend_set h B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   569
apply (cut_tac subset_refl [THEN stable_reachable, THEN project_leadsETo_D_lemma])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   570
apply (auto simp add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   571
 apply (erule leadsETo_mono [THEN [2] rev_subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   572
 apply (blast intro: extend_set_givenBy_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   573
apply (simp add: project_set_reachable_extend_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   574
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   575
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   576
lemma extending_leadsETo: 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   577
     "(ALL G. extend h F ok G --> G : preserves (v o f))  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   578
      ==> extending (%G. UNIV) h F  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   579
                (extend_set h A leadsTo[givenBy (v o f)] extend_set h B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   580
                (A leadsTo[givenBy v] B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   581
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   582
apply (auto simp add: project_leadsETo_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   583
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   584
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   585
lemma extending_LeadsETo: 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   586
     "(ALL G. extend h F ok G --> G : preserves (v o f))  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   587
      ==> extending (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   588
                (extend_set h A LeadsTo[givenBy (v o f)] extend_set h B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   589
                (A LeadsTo[givenBy v]  B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   590
apply (unfold extending_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   591
apply (blast intro: project_LeadsETo_D)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   592
done
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   593
*)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   594
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   595
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   596
(*** leadsETo in the precondition ***)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   597
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   598
(*Lemma for the Trans case*)
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   599
lemma pli_lemma:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   600
     "[| extend h F\<squnion>G : stable C;     
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   601
         F\<squnion>project h C G     
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   602
           : project_set h C Int project_set h A leadsTo project_set h B |]  
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   603
      ==> F\<squnion>project h C G     
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   604
            : project_set h C Int project_set h A leadsTo     
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   605
              project_set h C Int project_set h B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   606
apply (rule psp_stable2 [THEN leadsTo_weaken_L])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   607
apply (auto simp add: project_stable_project_set extend_stable_project_set)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   608
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   609
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   610
lemma project_leadsETo_I_lemma:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   611
     "[| extend h F\<squnion>G : stable C;   
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   612
         extend h F\<squnion>G :  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   613
           (C Int A) leadsTo[(%D. C Int D)`givenBy f]  B |]   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   614
  ==> F\<squnion>project h C G   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   615
    : (project_set h C Int project_set h (C Int A)) leadsTo (project_set h B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   616
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   617
  prefer 3
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   618
  apply (simp only: Int_UN_distrib project_set_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   619
  apply (blast intro: leadsTo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   620
 prefer 2 apply (blast intro: leadsTo_Trans pli_lemma)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   621
apply (simp add: givenBy_eq_extend_set)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   622
apply (rule leadsTo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   623
apply (blast intro: ensures_extend_set_imp_project_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   624
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   625
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   626
lemma project_leadsETo_I:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   627
     "extend h F\<squnion>G : (extend_set h A) leadsTo[givenBy f] (extend_set h B)
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   628
      ==> F\<squnion>project h UNIV G : A leadsTo B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   629
apply (rule project_leadsETo_I_lemma [THEN leadsTo_weaken], auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   630
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   631
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   632
lemma project_LeadsETo_I:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   633
     "extend h F\<squnion>G : (extend_set h A) LeadsTo[givenBy f] (extend_set h B) 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   634
      ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G   
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   635
           : A LeadsTo B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   636
apply (simp (no_asm_use) add: LeadsTo_def LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   637
apply (rule project_leadsETo_I_lemma [THEN leadsTo_weaken])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   638
apply (auto simp add: project_set_reachable_extend_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   639
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   640
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   641
lemma projecting_leadsTo: 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   642
     "projecting (%G. UNIV) h F  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   643
                 (extend_set h A leadsTo[givenBy f] extend_set h B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   644
                 (A leadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   645
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   646
apply (force dest: project_leadsETo_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   647
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   648
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   649
lemma projecting_LeadsTo: 
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   650
     "projecting (%G. reachable (extend h F\<squnion>G)) h F  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   651
                 (extend_set h A LeadsTo[givenBy f] extend_set h B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   652
                 (A LeadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   653
apply (unfold projecting_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   654
apply (force dest: project_LeadsETo_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   655
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   656
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   657
end
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   658
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   659
end