src/HOL/UNITY/PPROD.thy
author paulson
Fri, 24 Jan 2003 18:13:59 +0100
changeset 13786 ab8f39f48a6f
parent 8251 9be357df93d4
child 13798 4c1a53627500
permissions -rw-r--r--
More conversion of UNITY to Isar new-style theories
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/PPROD.thy
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     2
    ID:         $Id$
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     5
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6842
diff changeset
     6
Abstraction over replicated components (PLam)
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6842
diff changeset
     7
General products of programs (Pi operation)
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
     8
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
     9
Some dead wood here!
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    10
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    11
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    12
theory PPROD = Lift_prog:
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    13
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    14
constdefs
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    15
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7188
diff changeset
    16
  PLam  :: "[nat set, nat => ('b * ((nat=>'b) * 'c)) program]
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7188
diff changeset
    17
            => ((nat=>'b) * 'c) program"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7188
diff changeset
    18
    "PLam I F == JN i:I. lift i (F i)"
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    19
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    20
syntax
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    21
  "@PLam" :: "[pttrn, nat set, 'b set] => (nat => 'b) set"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    22
              ("(3plam _:_./ _)" 10)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    23
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    24
translations
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    25
  "plam x:A. B"   == "PLam A (%x. B)"
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    26
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    27
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    28
(*** Basic properties ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    29
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    30
lemma Init_PLam: "Init (PLam I F) = (INT i:I. lift_set i (Init (F i)))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    31
apply (simp (no_asm) add: PLam_def lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    32
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    33
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    34
declare Init_PLam [simp]
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    35
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    36
lemma PLam_empty: "PLam {} F = SKIP"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    37
apply (simp (no_asm) add: PLam_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    38
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    39
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    40
lemma PLam_SKIP: "(plam i: I. SKIP) = SKIP"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    41
apply (simp (no_asm) add: PLam_def lift_SKIP JN_constant)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    42
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    43
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    44
declare PLam_SKIP [simp] PLam_empty [simp]
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    45
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    46
lemma PLam_insert: "PLam (insert i I) F = (lift i (F i)) Join (PLam I F)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    47
by (unfold PLam_def, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    48
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    49
lemma PLam_component_iff: "((PLam I F) <= H) = (ALL i: I. lift i (F i) <= H)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    50
apply (simp (no_asm) add: PLam_def JN_component_iff)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    51
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    52
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    53
lemma component_PLam: "i : I ==> lift i (F i) <= (PLam I F)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    54
apply (unfold PLam_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    55
(*blast_tac doesn't use HO unification*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    56
apply (fast intro: component_JN)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    57
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    58
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    59
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    60
(** Safety & Progress: but are they used anywhere? **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    61
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    62
lemma PLam_constrains: "[| i : I;  ALL j. F j : preserves snd |] ==>   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    63
      (PLam I F : (lift_set i (A <*> UNIV)) co  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    64
                  (lift_set i (B <*> UNIV)))  =   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    65
      (F i : (A <*> UNIV) co (B <*> UNIV))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    66
apply (simp (no_asm_simp) add: PLam_def JN_constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    67
apply (subst insert_Diff [symmetric], assumption)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    68
apply (simp (no_asm_simp) add: lift_constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    69
apply (blast intro: constrains_imp_lift_constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    70
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    71
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    72
lemma PLam_stable: "[| i : I;  ALL j. F j : preserves snd |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    73
      ==> (PLam I F : stable (lift_set i (A <*> UNIV))) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    74
          (F i : stable (A <*> UNIV))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    75
apply (simp (no_asm_simp) add: stable_def PLam_constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    76
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    77
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    78
lemma PLam_transient: "i : I ==>  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    79
    PLam I F : transient A = (EX i:I. lift i (F i) : transient A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    80
apply (simp (no_asm_simp) add: JN_transient PLam_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    81
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    82
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    83
(*This holds because the F j cannot change (lift_set i)*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    84
lemma PLam_ensures: "[| i : I;  F i : (A <*> UNIV) ensures (B <*> UNIV);   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    85
         ALL j. F j : preserves snd |] ==>   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    86
      PLam I F : lift_set i (A <*> UNIV) ensures lift_set i (B <*> UNIV)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    87
apply (auto simp add: ensures_def PLam_constrains PLam_transient lift_transient_eq_disj lift_set_Un_distrib [symmetric] lift_set_Diff_distrib [symmetric] Times_Un_distrib1 [symmetric] Times_Diff_distrib1 [symmetric])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    88
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    89
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    90
lemma PLam_leadsTo_Basis: "[| i : I;   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    91
         F i : ((A <*> UNIV) - (B <*> UNIV)) co  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    92
               ((A <*> UNIV) Un (B <*> UNIV));   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    93
         F i : transient ((A <*> UNIV) - (B <*> UNIV));   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    94
         ALL j. F j : preserves snd |] ==>   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    95
      PLam I F : lift_set i (A <*> UNIV) leadsTo lift_set i (B <*> UNIV)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    96
by (rule PLam_ensures [THEN leadsTo_Basis], rule_tac [2] ensuresI)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    97
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    98
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
    99
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   100
(** invariant **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   101
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   102
lemma invariant_imp_PLam_invariant: "[| F i : invariant (A <*> UNIV);  i : I;   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   103
         ALL j. F j : preserves snd |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   104
      ==> PLam I F : invariant (lift_set i (A <*> UNIV))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   105
by (auto simp add: PLam_stable invariant_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   106
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   107
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   108
lemma PLam_preserves_fst [simp]:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   109
     "ALL j. F j : preserves snd  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   110
      ==> (PLam I F : preserves (v o sub j o fst)) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   111
          (if j: I then F j : preserves (v o fst) else True)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   112
by (simp (no_asm_simp) add: PLam_def lift_preserves_sub)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   113
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   114
lemma PLam_preserves_snd [simp,intro]:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   115
     "ALL j. F j : preserves snd ==> PLam I F : preserves snd"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   116
by (simp (no_asm_simp) add: PLam_def lift_preserves_snd_I)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   117
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   118
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   119
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   120
(*** guarantees properties ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   121
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   122
(*This rule looks unsatisfactory because it refers to "lift".  One must use
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   123
  lift_guarantees_eq_lift_inv to rewrite the first subgoal and
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   124
  something like lift_preserves_sub to rewrite the third.  However there's
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   125
  no obvious way to alternative for the third premise.*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   126
lemma guarantees_PLam_I: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   127
    "[| lift i (F i): X guarantees Y;  i : I;   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   128
        OK I (%i. lift i (F i)) |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   129
     ==> (PLam I F) : X guarantees Y"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   130
apply (unfold PLam_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   131
apply (simp (no_asm_simp) add: guarantees_JN_I)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   132
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   133
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   134
lemma Allowed_PLam [simp]:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   135
     "Allowed (PLam I F) = (INT i:I. lift i ` Allowed(F i))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   136
by (simp (no_asm) add: PLam_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   137
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   138
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   139
lemma PLam_preserves [simp]:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   140
     "(PLam I F) : preserves v = (ALL i:I. F i : preserves (v o lift_map i))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   141
by (simp (no_asm) add: PLam_def lift_def rename_preserves)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   142
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   143
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   144
(**UNUSED
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   145
    (*The f0 premise ensures that the product is well-defined.*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   146
    lemma PLam_invariant_imp_invariant: "[| PLam I F : invariant (lift_set i A);  i : I;   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   147
             f0: Init (PLam I F) |] ==> F i : invariant A"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   148
    apply (auto simp add: invariant_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   149
    apply (drule_tac c = "f0 (i:=x) " in subsetD)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   150
    apply auto
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   151
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   152
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   153
    lemma PLam_invariant: "[| i : I;  f0: Init (PLam I F) |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   154
          ==> (PLam I F : invariant (lift_set i A)) = (F i : invariant A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   155
    apply (blast intro: invariant_imp_PLam_invariant PLam_invariant_imp_invariant)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   156
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   157
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   158
    (*The f0 premise isn't needed if F is a constant program because then
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   159
      we get an initial state by replicating that of F*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   160
    lemma reachable_PLam: "i : I  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   161
          ==> ((plam x:I. F) : invariant (lift_set i A)) = (F : invariant A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   162
    apply (auto simp add: invariant_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   163
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   164
**)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   165
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   166
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   167
(**UNUSED
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   168
    (** Reachability **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   169
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   170
    Goal "[| f : reachable (PLam I F);  i : I |] ==> f i : reachable (F i)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   171
    apply (erule reachable.induct)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   172
    apply (auto intro: reachable.intrs)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   173
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   174
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   175
    (*Result to justify a re-organization of this file*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   176
    lemma ??unknown??: "{f. ALL i:I. f i : R i} = (INT i:I. lift_set i (R i))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   177
    apply auto
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   178
    result()
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   179
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   180
    lemma reachable_PLam_subset1: "reachable (PLam I F) <= (INT i:I. lift_set i (reachable (F i)))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   181
    apply (force dest!: reachable_PLam)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   182
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   183
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   184
    (*simplify using reachable_lift??*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   185
    lemma reachable_lift_Join_PLam [rule_format (no_asm)]: "[| i ~: I;  A : reachable (F i) |]      
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   186
       ==> ALL f. f : reachable (PLam I F)       
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   187
                  --> f(i:=A) : reachable (lift i (F i) Join PLam I F)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   188
    apply (erule reachable.induct)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   189
    apply (ALLGOALS Clarify_tac)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   190
    apply (erule reachable.induct)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   191
    (*Init, Init case*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   192
    apply (force intro: reachable.intrs)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   193
    (*Init of F, action of PLam F case*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   194
    apply (rule_tac act = act in reachable.Acts)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   195
    apply force
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   196
    apply assumption
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   197
    apply (force intro: ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   198
    (*induction over the 2nd "reachable" assumption*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   199
    apply (erule_tac xa = f in reachable.induct)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   200
    (*Init of PLam F, action of F case*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   201
    apply (rule_tac act = "lift_act i act" in reachable.Acts)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   202
    apply force
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   203
    apply (force intro: reachable.Init)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   204
    apply (force intro: ext simp add: lift_act_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   205
    (*last case: an action of PLam I F*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   206
    apply (rule_tac act = acta in reachable.Acts)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   207
    apply force
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   208
    apply assumption
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   209
    apply (force intro: ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   210
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   211
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   212
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   213
    (*The index set must be finite: otherwise infinitely many copies of F can
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   214
      perform actions, and PLam can never catch up in finite time.*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   215
    lemma reachable_PLam_subset2: "finite I  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   216
          ==> (INT i:I. lift_set i (reachable (F i))) <= reachable (PLam I F)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   217
    apply (erule finite_induct)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   218
    apply (simp (no_asm))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   219
    apply (force dest: reachable_lift_Join_PLam simp add: PLam_insert)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   220
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   221
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   222
    lemma reachable_PLam_eq: "finite I ==>  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   223
          reachable (PLam I F) = (INT i:I. lift_set i (reachable (F i)))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   224
    apply (REPEAT_FIRST (ares_tac [equalityI, reachable_PLam_subset1, reachable_PLam_subset2]))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   225
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   226
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   227
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   228
    (** Co **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   229
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   230
    lemma Constrains_imp_PLam_Constrains: "[| F i : A Co B;  i: I;  finite I |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   231
          ==> PLam I F : (lift_set i A) Co (lift_set i B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   232
    apply (auto simp add: Constrains_def Collect_conj_eq [symmetric] reachable_PLam_eq)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   233
    apply (auto simp add: constrains_def PLam_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   234
    apply (REPEAT (blast intro: reachable.intrs))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   235
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   236
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   237
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   238
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   239
    lemma PLam_Constrains: "[| i: I;  finite I;  f0: Init (PLam I F) |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   240
          ==> (PLam I F : (lift_set i A) Co (lift_set i B)) =   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   241
              (F i : A Co B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   242
    apply (blast intro: Constrains_imp_PLam_Constrains PLam_Constrains_imp_Constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   243
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   244
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   245
    lemma PLam_Stable: "[| i: I;  finite I;  f0: Init (PLam I F) |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   246
          ==> (PLam I F : Stable (lift_set i A)) = (F i : Stable A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   247
    apply (simp (no_asm_simp) del: Init_PLam add: Stable_def PLam_Constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   248
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   249
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   250
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   251
    (** const_PLam (no dependence on i) doesn't require the f0 premise **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   252
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   253
    lemma const_PLam_Constrains: "[| i: I;  finite I |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   254
          ==> ((plam x:I. F) : (lift_set i A) Co (lift_set i B)) =   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   255
              (F : A Co B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   256
    apply (blast intro: Constrains_imp_PLam_Constrains const_PLam_Constrains_imp_Constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   257
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   258
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   259
    lemma const_PLam_Stable: "[| i: I;  finite I |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   260
          ==> ((plam x:I. F) : Stable (lift_set i A)) = (F : Stable A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   261
    apply (simp (no_asm_simp) add: Stable_def const_PLam_Constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   262
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   263
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   264
    lemma const_PLam_Increasing: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   265
	 "[| i: I;  finite I |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   266
          ==> ((plam x:I. F) : Increasing (f o sub i)) = (F : Increasing f)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   267
    apply (unfold Increasing_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   268
    apply (subgoal_tac "ALL z. {s. z <= (f o sub i) s} = lift_set i {s. z <= f s}")
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   269
    apply (asm_simp_tac (simpset () add: lift_set_sub) 2)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   270
    apply (simp add: finite_lessThan const_PLam_Stable)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   271
    done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   272
**)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   273
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 8251
diff changeset
   274
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   275
end