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