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