src/HOL/UNITY/ELT.thy
author berghofe
Wed, 11 Jul 2007 11:46:44 +0200
changeset 23767 7272a839ccd9
parent 16417 9bc16273c2d4
child 26349 7f5a2f6d9119
permissions -rw-r--r--
Adapted to new inductive definition package.
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
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13836
diff changeset
    27
theory ELT imports Project begin
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    28
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    29
inductive_set
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    30
  (*LEADS-TO constant for the inductive definition*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    31
  elt :: "['a set set, 'a program] => ('a set * 'a set) set"
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    32
  for CC :: "'a set set" and F :: "'a program"
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    33
 where
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    34
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    35
   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
    36
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    37
 | 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
    38
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    39
 | 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
    40
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    41
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    42
constdefs
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
    43
  
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
    44
  (*the set of all sets determined by f alone*)
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
    45
  givenBy :: "['a => 'b] => 'a set set"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10293
diff changeset
    46
    "givenBy f == range (%B. f-` B)"
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    47
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    48
  (*visible version of the LEADS-TO relation*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    49
  leadsETo :: "['a set, 'a set set, 'a set] => 'a program set"
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    50
                                        ("(3_/ leadsTo[_]/ _)" [80,0,80] 80)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    51
    "leadsETo A CC B == {F. (A,B) : elt CC F}"
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    52
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)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    55
    "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)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   105
apply (tactic "deepen_tac (set_cs addSIs [equalityI]) 0 1")
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   106
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   107
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
(** Standard leadsTo rules **)
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
lemma leadsETo_Basis [intro]: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   112
     "[| 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
   113
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   114
apply (blast intro: elt.Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   115
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   116
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   117
lemma leadsETo_Trans: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   118
     "[| 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
   119
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   120
apply (blast intro: elt.Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   121
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   122
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
(*Useful with cancellation, disjunction*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   125
lemma leadsETo_Un_duplicate:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   126
     "F : A leadsTo[CC] (A' Un A') ==> F : A leadsTo[CC] A'"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   127
by (simp add: Un_ac)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   128
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   129
lemma leadsETo_Un_duplicate2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   130
     "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
   131
by (simp add: Un_ac)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   132
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   133
(*The Union introduction rule as we should have liked to state it*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   134
lemma leadsETo_Union:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   135
    "(!!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
   136
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   137
apply (blast intro: elt.Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   138
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   139
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   140
lemma leadsETo_UN:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   141
    "(!!i. i : I ==> F : (A i) leadsTo[CC] B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   142
     ==> F : (UN i:I. A i) leadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   143
apply (subst Union_image_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   144
apply (blast intro: leadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   145
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   146
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   147
(*The INDUCTION rule as we should have liked to state it*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   148
lemma leadsETo_induct:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   149
  "[| F : za leadsTo[CC] zb;   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   150
      !!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
   151
      !!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
   152
               ==> P A C;  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   153
      !!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
   154
   |] ==> P za zb"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   155
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   156
apply (drule CollectD) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   157
apply (erule elt.induct, blast+)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   158
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   159
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
(** New facts involving leadsETo **)
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
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
   164
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   165
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   166
prefer 3 apply (blast intro: leadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   167
prefer 2 apply (blast intro: leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   168
apply (blast intro: leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   169
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   170
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   171
lemma leadsETo_Trans_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   172
     "[| F : A leadsTo[CC] B;  F : B leadsTo[DD] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   173
      ==> F : A leadsTo[CC Un DD] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   174
by (blast intro: leadsETo_mono [THEN subsetD] leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   175
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   176
lemma leadsETo_Union_Int:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   177
 "(!!A. A : S ==> F : (A Int C) leadsTo[CC] B) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   178
  ==> F : (Union S Int C) leadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   179
apply (unfold leadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   180
apply (simp only: Int_Union_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   181
apply (blast intro: elt.Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   182
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   183
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   184
(*Binary union introduction rule*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   185
lemma leadsETo_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   186
     "[| F : A leadsTo[CC] C; F : B leadsTo[CC] C |] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   187
      ==> F : (A Un B) leadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   188
apply (subst Un_eq_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   189
apply (blast intro: leadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   190
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   191
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   192
lemma single_leadsETo_I:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   193
     "(!!x. x : A ==> F : {x} leadsTo[CC] B) ==> F : A leadsTo[CC] B"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   194
by (subst UN_singleton [symmetric], rule leadsETo_UN, blast)
13790
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   197
lemma subset_imp_leadsETo: "A<=B ==> F : A leadsTo[CC] B"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   198
by (simp add: subset_imp_ensures [THEN leadsETo_Basis] 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   199
              Diff_eq_empty_iff [THEN iffD2])
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   200
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   201
lemmas empty_leadsETo = empty_subsetI [THEN subset_imp_leadsETo, simp]
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   205
(** Weakening laws **)
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
lemma leadsETo_weaken_R:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   208
     "[| F : A leadsTo[CC] A';  A'<=B' |] ==> F : A leadsTo[CC] B'"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   209
by (blast intro: subset_imp_leadsETo leadsETo_Trans)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   210
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   211
lemma leadsETo_weaken_L [rule_format]:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   212
     "[| F : A leadsTo[CC] A'; B<=A |] ==> F : B leadsTo[CC] A'"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   213
by (blast intro: leadsETo_Trans subset_imp_leadsETo)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   214
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   215
(*Distributes over binary unions*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   216
lemma leadsETo_Un_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   217
     "F : (A Un B) leadsTo[CC] C  =   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   218
      (F : A leadsTo[CC] C & F : B leadsTo[CC] C)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   219
by (blast intro: leadsETo_Un leadsETo_weaken_L)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   220
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   221
lemma leadsETo_UN_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   222
     "F : (UN i:I. A i) leadsTo[CC] B  =   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   223
      (ALL i : I. F : (A i) leadsTo[CC] B)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   224
by (blast intro: leadsETo_UN leadsETo_weaken_L)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   225
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   226
lemma leadsETo_Union_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   227
     "F : (Union S) leadsTo[CC] B  =  (ALL A : S. F : A leadsTo[CC] B)"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   228
by (blast intro: leadsETo_Union leadsETo_weaken_L)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   229
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   230
lemma leadsETo_weaken:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   231
     "[| F : A leadsTo[CC'] A'; B<=A; A'<=B';  CC' <= CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   232
      ==> F : B leadsTo[CC] B'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   233
apply (drule leadsETo_mono [THEN subsetD], assumption)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   234
apply (blast del: subsetCE 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   235
             intro: leadsETo_weaken_R leadsETo_weaken_L leadsETo_Trans)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   236
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   237
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   238
lemma leadsETo_givenBy:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   239
     "[| F : A leadsTo[CC] A';  CC <= givenBy v |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   240
      ==> F : A leadsTo[givenBy v] A'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   241
by (blast intro: empty_mem_givenBy leadsETo_weaken)
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   244
(*Set difference*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   245
lemma leadsETo_Diff:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   246
     "[| F : (A-B) leadsTo[CC] C; F : B leadsTo[CC] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   247
      ==> F : A leadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   248
by (blast intro: leadsETo_Un leadsETo_weaken)
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   251
(*Binary union version*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   252
lemma leadsETo_Un_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   253
     "[| F : A leadsTo[CC] A';  F : B leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   254
      ==> F : (A Un B) leadsTo[CC] (A' Un B')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   255
by (blast intro: leadsETo_Un leadsETo_weaken_R)
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   258
(** The cancellation law **)
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
lemma leadsETo_cancel2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   261
     "[| F : A leadsTo[CC] (A' Un B); F : B leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   262
      ==> F : A leadsTo[CC] (A' Un B')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   263
by (blast intro: leadsETo_Un_Un subset_imp_leadsETo leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   264
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   265
lemma leadsETo_cancel1:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   266
     "[| F : A leadsTo[CC] (B Un A'); F : B leadsTo[CC] B' |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   267
    ==> F : A leadsTo[CC] (B' Un A')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   268
apply (simp add: Un_commute)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   269
apply (blast intro!: leadsETo_cancel2)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   270
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   271
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   272
lemma leadsETo_cancel_Diff1:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   273
     "[| 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
   274
    ==> F : A leadsTo[CC] (B' Un A')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   275
apply (rule leadsETo_cancel1)
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   276
 prefer 2 apply assumption
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13798
diff changeset
   277
apply simp_all
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   278
done
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   281
(** PSP: Progress-Safety-Progress **)
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
(*Special case of PSP: Misra's "stable conjunction"*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   284
lemma e_psp_stable: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   285
   "[| 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
   286
    ==> F : (A Int B) leadsTo[CC] (A' Int B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   287
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   288
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   289
prefer 3 apply (blast intro: leadsETo_Union_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   290
prefer 2 apply (blast intro: leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   291
apply (rule leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   292
prefer 2 apply (force simp add: Diff_Int_distrib2 [symmetric])
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   293
apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   294
                 Int_Un_distrib2 [symmetric])
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   295
apply (blast intro: transient_strengthen constrains_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   296
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   297
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   298
lemma e_psp_stable2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   299
     "[| 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
   300
      ==> F : (B Int A) leadsTo[CC] (B Int A')"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   301
by (simp (no_asm_simp) add: e_psp_stable Int_ac)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   302
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   303
lemma e_psp:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   304
     "[| F : A leadsTo[CC] A'; F : B co B';   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   305
         ALL C:CC. C Int B Int B' : CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   306
      ==> F : (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   307
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   308
prefer 3 apply (blast intro: leadsETo_Union_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   309
(*Transitivity case has a delicate argument involving "cancellation"*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   310
apply (rule_tac [2] leadsETo_Un_duplicate2)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   311
apply (erule_tac [2] leadsETo_cancel_Diff1)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   312
prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   313
 apply (simp add: Int_Diff Diff_triv)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   314
 apply (blast intro: leadsETo_weaken_L dest: constrains_imp_subset)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   315
(*Basis case*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   316
apply (rule leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   317
apply (blast intro: psp_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   318
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
   319
apply auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   320
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   321
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   322
lemma e_psp2:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   323
     "[| F : A leadsTo[CC] A'; F : B co B';   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   324
         ALL C:CC. C Int B Int B' : CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   325
      ==> F : (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   326
by (simp add: e_psp Int_ac)
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   329
(*** Special properties involving the parameter [CC] ***)
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
(*??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
   332
lemma gen_leadsETo_imp_Join_leadsETo:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   333
     "[| F: (A leadsTo[givenBy v] B);  G : preserves v;   
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   334
         F\<squnion>G : stable C |]  
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   335
      ==> 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
   336
apply (erule leadsETo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   337
  prefer 3
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   338
  apply (subst Int_Union) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   339
  apply (blast intro: leadsETo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   340
prefer 2
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   341
 apply (blast intro: e_psp_stable2 [THEN leadsETo_weaken_L] leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   342
apply (rule leadsETo_Basis)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   343
apply (auto simp add: Diff_eq_empty_iff [THEN iffD2] 
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   344
                      Int_Diff ensures_def givenBy_eq_Collect Join_transient)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   345
prefer 3 apply (blast intro: transient_strengthen)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   346
apply (drule_tac [2] P1 = P in preserves_subset_stable [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   347
apply (drule_tac P1 = P in preserves_subset_stable [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   348
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   349
apply (blast intro: constrains_Int [THEN constrains_weaken])+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   350
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   351
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   352
(**** Relationship with traditional "leadsTo", strong & weak ****)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   353
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   354
(** strong **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   355
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   356
lemma leadsETo_subset_leadsTo: "(A leadsTo[CC] B) <= (A leadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   357
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   358
apply (erule leadsETo_induct)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   359
  prefer 3 apply (blast intro: leadsTo_Union)
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   360
 prefer 2 apply (blast intro: leadsTo_Trans, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   361
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   362
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   363
lemma leadsETo_UNIV_eq_leadsTo: "(A leadsTo[UNIV] B) = (A leadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   364
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   365
apply (erule leadsETo_subset_leadsTo [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   366
(*right-to-left case*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   367
apply (erule leadsTo_induct)
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   368
  prefer 3 apply (blast intro: leadsETo_Union)
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   369
 prefer 2 apply (blast intro: leadsETo_Trans, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   370
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   371
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   372
(**** weak ****)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   373
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   374
lemma LeadsETo_eq_leadsETo: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   375
     "A LeadsTo[CC] B =  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   376
        {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
   377
        (reachable F Int B)}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   378
apply (unfold LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   379
apply (blast dest: e_psp_stable2 intro: leadsETo_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   380
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   381
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   382
(*** Introduction rules: Basis, Trans, Union ***)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   383
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   384
lemma LeadsETo_Trans:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   385
     "[| F : A LeadsTo[CC] B;  F : B LeadsTo[CC] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   386
      ==> F : A LeadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   387
apply (simp add: LeadsETo_eq_leadsETo)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   388
apply (blast intro: leadsETo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   389
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   390
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   391
lemma LeadsETo_Union:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   392
     "(!!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
   393
apply (simp add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   394
apply (subst Int_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   395
apply (blast intro: leadsETo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   396
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   397
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   398
lemma LeadsETo_UN:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   399
     "(!!i. i : I ==> F : (A i) LeadsTo[CC] B)  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   400
      ==> F : (UN i:I. A i) LeadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   401
apply (simp only: Union_image_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   402
apply (blast intro: LeadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   403
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   404
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   405
(*Binary union introduction rule*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   406
lemma LeadsETo_Un:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   407
     "[| F : A LeadsTo[CC] C; F : B LeadsTo[CC] C |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   408
      ==> F : (A Un B) LeadsTo[CC] C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   409
apply (subst Un_eq_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   410
apply (blast intro: LeadsETo_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   411
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   412
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   413
(*Lets us look at the starting state*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   414
lemma single_LeadsETo_I:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   415
     "(!!s. s : A ==> F : {s} LeadsTo[CC] B) ==> F : A LeadsTo[CC] B"
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   416
by (subst UN_singleton [symmetric], rule LeadsETo_UN, blast)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   417
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   418
lemma subset_imp_LeadsETo:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   419
     "A <= B ==> F : A LeadsTo[CC] B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   420
apply (simp (no_asm) add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   421
apply (blast intro: subset_imp_leadsETo)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   422
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   423
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   424
lemmas empty_LeadsETo = empty_subsetI [THEN subset_imp_LeadsETo, standard]
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   425
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   426
lemma LeadsETo_weaken_R [rule_format]:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   427
     "[| F : A LeadsTo[CC] A';  A' <= B' |] ==> F : A LeadsTo[CC] B'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   428
apply (simp (no_asm_use) add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   429
apply (blast intro: leadsETo_weaken_R)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   430
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   431
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   432
lemma LeadsETo_weaken_L [rule_format]:
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   433
     "[| F : A LeadsTo[CC] A';  B <= A |] ==> F : B LeadsTo[CC] A'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   434
apply (simp (no_asm_use) add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   435
apply (blast intro: leadsETo_weaken_L)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   436
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   437
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   438
lemma LeadsETo_weaken:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   439
     "[| F : A LeadsTo[CC'] A';    
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   440
         B <= A;  A' <= B';  CC' <= CC |]  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   441
      ==> F : B LeadsTo[CC] B'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   442
apply (simp (no_asm_use) add: LeadsETo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   443
apply (blast intro: leadsETo_weaken)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   444
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   445
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   446
lemma LeadsETo_subset_LeadsTo: "(A LeadsTo[CC] B) <= (A LeadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   447
apply (unfold LeadsETo_def LeadsTo_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   448
apply (blast intro: leadsETo_subset_leadsTo [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   449
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   450
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   451
(*Postcondition can be strengthened to (reachable F Int B) *)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   452
lemma reachable_ensures:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   453
     "F : A ensures B ==> F : (reachable F Int A) ensures B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   454
apply (rule stable_ensures_Int [THEN ensures_weaken_R], auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   455
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   456
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   457
lemma lel_lemma:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   458
     "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
   459
apply (erule leadsTo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   460
  apply (blast intro: reachable_ensures leadsETo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   461
 apply (blast dest: e_psp_stable2 intro: leadsETo_Trans leadsETo_weaken_L)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   462
apply (subst Int_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   463
apply (blast intro: leadsETo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   464
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   465
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   466
lemma LeadsETo_UNIV_eq_LeadsTo: "(A LeadsTo[UNIV] B) = (A LeadsTo B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   467
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   468
apply (erule LeadsETo_subset_LeadsTo [THEN subsetD])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   469
(*right-to-left case*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   470
apply (unfold LeadsETo_def LeadsTo_def)
13836
6d0392fc6dc5 restored some deleted lemmas
paulson
parents: 13819
diff changeset
   471
apply (blast intro: lel_lemma [THEN leadsETo_weaken])
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   472
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   473
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   474
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   475
(**** EXTEND/PROJECT PROPERTIES ****)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   476
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   477
lemma (in Extend) givenBy_o_eq_extend_set:
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
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)"
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   487
lemma (in Extend) extend_set_givenBy_I:
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   492
lemma (in Extend) leadsETo_imp_extend_leadsETo:
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)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   497
  apply (force intro: leadsETo_Basis subset_imp_ensures 
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*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   506
lemma (in Extend) Join_project_ensures_strong:
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
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   518
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   541
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   553
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   563
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   576
lemma (in Extend) extending_leadsETo: 
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   585
lemma (in Extend) extending_LeadsETo: 
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*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   599
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   610
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   626
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   632
lemma (in Extend) 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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   641
lemma (in Extend) projecting_leadsTo: 
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
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   649
lemma (in Extend) 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