src/HOL/UNITY/Lift_prog.thy
author paulson
Wed, 29 Jan 2003 16:34:51 +0100
changeset 13792 d1811693899c
parent 13790 8d7e9fce8c50
child 13798 4c1a53627500
permissions -rw-r--r--
converted more UNITY theories to new-style
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Lift_prog.thy
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     2
    ID:         $Id$
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     5
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
     6
lift_prog, etc: replication of components and arrays of processes. 
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     7
*)
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     8
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
     9
theory Lift_prog = Rename:
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    10
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    11
constdefs
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    12
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    13
  insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    14
    "insert_map i z f k == if k<i then f k
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    15
                           else if k=i then z
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10834
diff changeset
    16
                           else f(k - 1)"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    17
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    18
  delete_map :: "[nat, nat=>'b] => (nat=>'b)"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    19
    "delete_map i g k == if k<i then g k else g (Suc k)"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    20
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    21
  lift_map :: "[nat, 'b * ((nat=>'b) * 'c)] => (nat=>'b) * 'c"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    22
    "lift_map i == %(s,(f,uu)). (insert_map i s f, uu)"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    23
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    24
  drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    25
    "drop_map i == %(g, uu). (g i, (delete_map i g, uu))"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    26
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    27
  lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 8251
diff changeset
    28
    "lift_set i A == lift_map i ` A"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    29
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    30
  lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
    31
    "lift i == rename (lift_map i)"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    32
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    33
  (*simplifies the expression of specifications*)
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    34
  sub :: "['a, 'a=>'b] => 'b"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    35
    "sub == %i f. f i"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    36
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    37
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    38
declare insert_map_def [simp] delete_map_def [simp]
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    39
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    40
lemma insert_map_inverse: "delete_map i (insert_map i x f) = f"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    41
apply (rule ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    42
apply (simp (no_asm))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    43
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    44
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    45
lemma insert_map_delete_map_eq: "(insert_map i x (delete_map i g)) = g(i:=x)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    46
apply (rule ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    47
apply (auto split add: nat_diff_split)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    48
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    49
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    50
(*** Injectiveness proof ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    51
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    52
lemma insert_map_inject1: "(insert_map i x f) = (insert_map i y g) ==> x=y"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    53
apply (drule_tac x = i in fun_cong)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    54
apply (simp (no_asm_use))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    55
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    56
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    57
lemma insert_map_inject2: "(insert_map i x f) = (insert_map i y g) ==> f=g"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    58
apply (drule_tac f = "delete_map i" in arg_cong)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    59
apply (simp (no_asm_use) add: insert_map_inverse)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    60
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    61
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    62
lemma insert_map_inject':
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    63
     "(insert_map i x f) = (insert_map i y g) ==> x=y & f=g"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    64
by (blast dest: insert_map_inject1 insert_map_inject2)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    65
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    66
lemmas insert_map_inject = insert_map_inject' [THEN conjE, elim!]
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    67
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    68
(*The general case: we don't assume i=i'*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    69
lemma lift_map_eq_iff [iff]: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    70
     "(lift_map i (s,(f,uu)) = lift_map i' (s',(f',uu')))  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    71
      = (uu = uu' & insert_map i s f = insert_map i' s' f')"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    72
apply (unfold lift_map_def, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    73
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    74
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    75
(*The !!s allows the automatic splitting of the bound variable*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    76
lemma drop_map_lift_map_eq [simp]: "!!s. drop_map i (lift_map i s) = s"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    77
apply (unfold lift_map_def drop_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    78
apply (force intro: insert_map_inverse)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    79
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    80
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    81
lemma inj_lift_map: "inj (lift_map i)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    82
apply (unfold lift_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    83
apply (rule inj_onI, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    84
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    85
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    86
(*** Surjectiveness proof ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    87
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    88
lemma lift_map_drop_map_eq [simp]: "!!s. lift_map i (drop_map i s) = s"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    89
apply (unfold lift_map_def drop_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    90
apply (force simp add: insert_map_delete_map_eq)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    91
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    92
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    93
lemma drop_map_inject [dest!]: "(drop_map i s) = (drop_map i s') ==> s=s'"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    94
apply (drule_tac f = "lift_map i" in arg_cong)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    95
apply (simp (no_asm_use))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    96
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    97
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    98
lemma surj_lift_map: "surj (lift_map i)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    99
apply (rule surjI)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   100
apply (rule lift_map_drop_map_eq)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   101
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   102
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   103
lemma bij_lift_map [iff]: "bij (lift_map i)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   104
apply (simp (no_asm) add: bij_def inj_lift_map surj_lift_map)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   105
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   106
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   107
lemma inv_lift_map_eq [simp]: "inv (lift_map i) = drop_map i"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   108
by (rule inv_equality, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   109
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   110
lemma inv_drop_map_eq [simp]: "inv (drop_map i) = lift_map i"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   111
by (rule inv_equality, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   112
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   113
lemma bij_drop_map [iff]: "bij (drop_map i)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   114
by (simp del: inv_lift_map_eq add: inv_lift_map_eq [symmetric] bij_imp_bij_inv)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   115
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   116
(*sub's main property!*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   117
lemma sub_apply [simp]: "sub i f = f i"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   118
apply (simp (no_asm) add: sub_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   119
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   120
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   121
(*** lift_set ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   122
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   123
lemma lift_set_empty [simp]: "lift_set i {} = {}"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   124
by (unfold lift_set_def, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   125
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   126
lemma lift_set_iff: "(lift_map i x : lift_set i A) = (x : A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   127
apply (unfold lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   128
apply (rule inj_lift_map [THEN inj_image_mem_iff])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   129
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   130
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   131
(*Do we really need both this one and its predecessor?*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   132
lemma lift_set_iff2 [iff]:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   133
     "((f,uu) : lift_set i A) = ((f i, (delete_map i f, uu)) : A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   134
by (simp (no_asm_simp) add: lift_set_def mem_rename_set_iff drop_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   135
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   136
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   137
lemma lift_set_mono: "A<=B ==> lift_set i A <= lift_set i B"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   138
apply (unfold lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   139
apply (erule image_mono)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   140
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   141
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   142
lemma lift_set_Un_distrib: "lift_set i (A Un B) = lift_set i A Un lift_set i B"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   143
apply (unfold lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   144
apply (simp (no_asm_simp) add: image_Un)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   145
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   146
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   147
lemma lift_set_Diff_distrib: "lift_set i (A-B) = lift_set i A - lift_set i B"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   148
apply (unfold lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   149
apply (rule inj_lift_map [THEN image_set_diff])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   150
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   151
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   152
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   153
(*** the lattice operations ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   154
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   155
lemma bij_lift [iff]: "bij (lift i)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   156
apply (simp (no_asm) add: lift_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   157
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   158
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   159
lemma lift_SKIP [simp]: "lift i SKIP = SKIP"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   160
apply (unfold lift_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   161
apply (simp (no_asm_simp))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   162
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   163
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   164
lemma lift_Join [simp]: "lift i (F Join G) = lift i F Join lift i G"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   165
apply (unfold lift_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   166
apply (simp (no_asm_simp))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   167
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   168
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   169
lemma lift_JN [simp]: "lift j (JOIN I F) = (JN i:I. lift j (F i))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   170
apply (unfold lift_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   171
apply (simp (no_asm_simp))
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   172
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   173
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   174
(*** Safety: co, stable, invariant ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   175
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   176
lemma lift_constrains: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   177
     "(lift i F : (lift_set i A) co (lift_set i B)) = (F : A co B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   178
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   179
apply (simp (no_asm_simp) add: rename_constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   180
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   181
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   182
lemma lift_stable: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   183
     "(lift i F : stable (lift_set i A)) = (F : stable A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   184
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   185
apply (simp (no_asm_simp) add: rename_stable)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   186
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   187
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   188
lemma lift_invariant: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   189
     "(lift i F : invariant (lift_set i A)) = (F : invariant A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   190
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   191
apply (simp (no_asm_simp) add: rename_invariant)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   192
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   193
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   194
lemma lift_Constrains: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   195
     "(lift i F : (lift_set i A) Co (lift_set i B)) = (F : A Co B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   196
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   197
apply (simp (no_asm_simp) add: rename_Constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   198
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   199
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   200
lemma lift_Stable: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   201
     "(lift i F : Stable (lift_set i A)) = (F : Stable A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   202
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   203
apply (simp (no_asm_simp) add: rename_Stable)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   204
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   205
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   206
lemma lift_Always: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   207
     "(lift i F : Always (lift_set i A)) = (F : Always A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   208
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   209
apply (simp (no_asm_simp) add: rename_Always)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   210
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   211
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   212
(*** Progress: transient, ensures ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   213
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   214
lemma lift_transient: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   215
     "(lift i F : transient (lift_set i A)) = (F : transient A)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   216
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   217
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   218
apply (simp (no_asm_simp) add: rename_transient)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   219
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   220
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   221
lemma lift_ensures: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   222
     "(lift i F : (lift_set i A) ensures (lift_set i B)) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   223
      (F : A ensures B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   224
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   225
apply (simp (no_asm_simp) add: rename_ensures)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   226
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   227
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   228
lemma lift_leadsTo: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   229
     "(lift i F : (lift_set i A) leadsTo (lift_set i B)) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   230
      (F : A leadsTo B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   231
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   232
apply (simp (no_asm_simp) add: rename_leadsTo)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   233
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   234
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   235
lemma lift_LeadsTo: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   236
     "(lift i F : (lift_set i A) LeadsTo (lift_set i B)) =   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   237
      (F : A LeadsTo B)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   238
apply (unfold lift_def lift_set_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   239
apply (simp (no_asm_simp) add: rename_LeadsTo)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   240
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   241
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   242
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   243
(** guarantees **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   244
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   245
lemma lift_lift_guarantees_eq: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   246
     "(lift i F : (lift i ` X) guarantees (lift i ` Y)) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   247
      (F : X guarantees Y)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   248
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   249
apply (unfold lift_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   250
apply (subst bij_lift_map [THEN rename_rename_guarantees_eq, symmetric])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   251
apply (simp (no_asm_simp) add: o_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   252
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   253
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   254
lemma lift_guarantees_eq_lift_inv: "(lift i F : X guarantees Y) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   255
      (F : (rename (drop_map i) ` X) guarantees (rename (drop_map i) ` Y))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   256
by (simp add: bij_lift_map [THEN rename_guarantees_eq_rename_inv] lift_def)
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
   257
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
   258
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   259
(*** We devote an ENORMOUS effort to proving lift_transient_eq_disj, 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   260
     which is used only in TimerArray and perhaps isn't even essential
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   261
     there!
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   262
***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   263
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   264
(*To preserve snd means that the second component is there just to allow
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   265
  guarantees properties to be stated.  Converse fails, for lift i F can 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   266
  change function components other than i*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   267
lemma lift_preserves_snd_I: "F : preserves snd ==> lift i F : preserves snd"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   268
apply (drule_tac w1=snd in subset_preserves_o [THEN subsetD])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   269
apply (simp (no_asm_simp) add: lift_def rename_preserves)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   270
apply (simp (no_asm_use) add: lift_map_def o_def split_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   271
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   272
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   273
lemma delete_map_eqE':
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   274
     "(delete_map i g) = (delete_map i g') ==> EX x. g = g'(i:=x)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   275
apply (drule_tac f = "insert_map i (g i) " in arg_cong)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   276
apply (simp (no_asm_use) add: insert_map_delete_map_eq)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   277
apply (erule exI)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   278
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   279
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   280
lemmas delete_map_eqE = delete_map_eqE' [THEN exE, elim!]
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   281
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   282
lemma delete_map_neq_apply:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   283
     "[| delete_map j g = delete_map j g';  i~=j |] ==> g i = g' i"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   284
by force
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   285
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   286
(*A set of the form (A <*> UNIV) ignores the second (dummy) state component*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   287
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   288
lemma vimage_o_fst_eq [simp]: "(f o fst) -` A = (f-`A) <*> UNIV"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   289
by auto
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   290
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   291
lemma vimage_sub_eq_lift_set [simp]:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   292
     "(sub i -`A) <*> UNIV = lift_set i (A <*> UNIV)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   293
by auto
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   294
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   295
lemma mem_lift_act_iff [iff]: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   296
     "((s,s') : extend_act (%(x,u::unit). lift_map i x) act) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   297
      ((drop_map i s, drop_map i s') : act)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   298
apply (unfold extend_act_def, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   299
apply (rule bexI, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   300
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   301
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   302
lemma preserves_snd_lift_stable:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   303
     "[| F : preserves snd;  i~=j |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   304
      ==> lift j F : stable (lift_set i (A <*> UNIV))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   305
apply (auto simp add: lift_def lift_set_def stable_def constrains_def rename_def extend_def mem_rename_set_iff)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   306
apply (auto dest!: preserves_imp_eq simp add: lift_map_def drop_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   307
apply (drule_tac x = i in fun_cong, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   308
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   309
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   310
(*If i~=j then lift j F  does nothing to lift_set i, and the 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   311
  premise ensures A<=B.*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   312
lemma constrains_imp_lift_constrains:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   313
    "[| F i : (A <*> UNIV) co (B <*> UNIV);   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   314
        F j : preserves snd |]   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   315
     ==> lift j (F j) : (lift_set i (A <*> UNIV)) co (lift_set i (B <*> UNIV))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   316
apply (case_tac "i=j")
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   317
apply (simp add: lift_def lift_set_def rename_constrains)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   318
apply (erule preserves_snd_lift_stable[THEN stableD, THEN constrains_weaken_R], assumption)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   319
apply (erule constrains_imp_subset [THEN lift_set_mono])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   320
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   321
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   322
(** Lemmas for the transient theorem **)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   323
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   324
lemma insert_map_upd_same: "(insert_map i t f)(i := s) = insert_map i s f"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   325
by (rule ext, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   326
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   327
lemma insert_map_upd:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   328
     "(insert_map j t f)(i := s) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   329
      (if i=j then insert_map i s f  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   330
       else if i<j then insert_map j t (f(i:=s))  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   331
       else insert_map j t (f(i - Suc 0 := s)))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   332
apply (rule ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   333
apply (simp split add: nat_diff_split) 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   334
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   335
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   336
lemma insert_map_eq_diff:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   337
     "[| insert_map i s f = insert_map j t g;  i~=j |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   338
      ==> EX g'. insert_map i s' f = insert_map j t g'"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   339
apply (subst insert_map_upd_same [symmetric])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   340
apply (erule ssubst)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   341
apply (simp only: insert_map_upd if_False split: split_if, blast)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   342
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   343
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   344
lemma lift_map_eq_diff: 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   345
     "[| lift_map i (s,(f,uu)) = lift_map j (t,(g,vv));  i~=j |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   346
      ==> EX g'. lift_map i (s',(f,uu)) = lift_map j (t,(g',vv))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   347
apply (unfold lift_map_def, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   348
apply (blast dest: insert_map_eq_diff)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   349
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   350
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   351
lemma lift_transient_eq_disj:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   352
     "F : preserves snd  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   353
      ==> (lift i F : transient (lift_set j (A <*> UNIV))) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   354
          (i=j & F : transient (A <*> UNIV) | A={})"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   355
apply (case_tac "i=j")
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   356
apply (auto simp add: lift_transient)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   357
apply (auto simp add: lift_set_def lift_def transient_def rename_def extend_def Domain_extend_act)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   358
apply (drule subsetD, blast, auto)
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   359
apply (rename_tac s f uu s' f' uu')
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   360
apply (subgoal_tac "f'=f & uu'=uu")
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   361
 prefer 2 apply (force dest!: preserves_imp_eq, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   362
apply (drule sym)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   363
apply (drule subsetD)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   364
apply (rule ImageI)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   365
apply (erule bij_lift_map [THEN good_map_bij, 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   366
                           THEN Extend.intro, 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   367
                           THEN Extend.mem_extend_act_iff [THEN iffD2]], force)
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   368
apply (erule lift_map_eq_diff [THEN exE], auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   369
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   370
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   371
(*USELESS??*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   372
lemma lift_map_image_Times: "lift_map i ` (A <*> UNIV) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   373
      (UN s:A. UN f. {insert_map i s f}) <*> UNIV"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   374
apply (auto intro!: bexI image_eqI simp add: lift_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   375
apply (rule split_conv [symmetric])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   376
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   377
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   378
lemma lift_preserves_eq:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   379
     "(lift i F : preserves v) = (F : preserves (v o lift_map i))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   380
by (simp add: lift_def rename_preserves)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   381
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   382
(*A useful rewrite.  If o, sub have been rewritten out already then can also
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   383
  use it as   rewrite_rule [sub_def, o_def] lift_preserves_sub*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   384
lemma lift_preserves_sub:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   385
     "F : preserves snd  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   386
      ==> lift i F : preserves (v o sub j o fst) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   387
          (if i=j then F : preserves (v o fst) else True)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   388
apply (drule subset_preserves_o [THEN subsetD])
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   389
apply (simp add: lift_preserves_eq o_def drop_map_lift_map_eq)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   390
apply (simp cong del: if_weak_cong add: lift_map_def eq_commute split_def o_def, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   391
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   392
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   393
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   394
(*** Lemmas to handle function composition (o) more consistently ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   395
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   396
(*Lets us prove one version of a theorem and store others*)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   397
lemma o_equiv_assoc: "f o g = h ==> f' o f o g = f' o h"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   398
by (simp add: expand_fun_eq o_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   399
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   400
lemma o_equiv_apply: "f o g = h ==> ALL x. f(g x) = h x"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   401
by (simp add: expand_fun_eq o_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   402
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   403
lemma fst_o_lift_map: "sub i o fst o lift_map i = fst"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   404
apply (rule ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   405
apply (auto simp add: o_def lift_map_def sub_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   406
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   407
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   408
lemma snd_o_lift_map: "snd o lift_map i = snd o snd"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   409
apply (rule ext)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   410
apply (auto simp add: o_def lift_map_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   411
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   412
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   413
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   414
(*** More lemmas about extend and project 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   415
     They could be moved to {Extend,Project}.ML, but DON'T need the locale ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   416
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   417
lemma extend_act_extend_act: "extend_act h' (extend_act h act) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   418
      extend_act (%(x,(y,y')). h'(h(x,y),y')) act"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   419
apply (auto elim!: rev_bexI simp add: extend_act_def, blast) 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   420
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   421
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   422
lemma project_act_project_act: "project_act h (project_act h' act) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   423
      project_act (%(x,(y,y')). h'(h(x,y),y')) act"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   424
by (auto elim!: rev_bexI simp add: project_act_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   425
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   426
lemma project_act_extend_act:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   427
     "project_act h (extend_act h' act) =  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   428
        {(x,x'). EX s s' y y' z. (s,s') : act &  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   429
                 h(x,y) = h'(s,z) & h(x',y') = h'(s',z)}"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   430
by (simp add: extend_act_def project_act_def, blast)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   431
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   432
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   433
(*** OK and "lift" ***)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   434
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   435
lemma act_in_UNION_preserves_fst:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   436
     "act <= {(x,x'). fst x = fst x'} ==> act : UNION (preserves fst) Acts"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   437
apply (rule_tac a = "mk_program (UNIV,{act},UNIV) " in UN_I)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   438
apply (auto simp add: preserves_def stable_def constrains_def)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   439
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   440
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   441
lemma UNION_OK_lift_I:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   442
     "[| ALL i:I. F i : preserves snd;   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   443
         ALL i:I. UNION (preserves fst) Acts <= AllowedActs (F i) |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   444
      ==> OK I (%i. lift i (F i))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   445
apply (auto simp add: OK_def lift_def rename_def Extend.Acts_extend)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   446
apply (simp (no_asm) add: Extend.AllowedActs_extend project_act_extend_act)
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   447
apply (rename_tac "act")
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   448
apply (subgoal_tac
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   449
       "{(x, x'). \<exists>s f u s' f' u'. 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   450
                    ((s, f, u), s', f', u') : act & 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   451
                    lift_map j x = lift_map i (s, f, u) & 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   452
                    lift_map j x' = lift_map i (s', f', u') } 
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   453
                <= { (x,x') . fst x = fst x'}")
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   454
apply (blast intro: act_in_UNION_preserves_fst, clarify)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   455
apply (drule_tac x = j in fun_cong)+
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   456
apply (drule_tac x = i in bspec, assumption)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   457
apply (frule preserves_imp_eq, auto)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   458
done
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   459
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   460
lemma OK_lift_I:
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   461
     "[| ALL i:I. F i : preserves snd;   
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   462
         ALL i:I. preserves fst <= Allowed (F i) |]  
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   463
      ==> OK I (%i. lift i (F i))"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   464
by (simp add: safety_prop_AllowedActs_iff_Allowed UNION_OK_lift_I)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   465
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   466
lemma Allowed_lift [simp]: "Allowed (lift i F) = lift i ` (Allowed F)"
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   467
by (simp add: lift_def Allowed_rename)
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   468
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   469
lemma lift_image_preserves:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   470
     "lift i ` preserves v = preserves (v o drop_map i)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 13786
diff changeset
   471
by (simp (no_asm) add: rename_image_preserves lift_def inv_lift_map_eq)
13786
ab8f39f48a6f More conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   472
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
   473
end