src/HOL/UNITY/Rename.thy
author wenzelm
Thu, 13 Mar 2025 11:19:27 +0100
changeset 82266 cca7113dcafc
parent 63807 5f77017055a3
permissions -rw-r--r--
Added tag Isabelle2025 for changeset 4b875a4c83b0
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8256
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Rename.thy
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
     3
    Copyright   2000  University of Cambridge
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
     4
*)
8256
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
     5
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
     6
section\<open>Renaming of State Sets\<close>
8256
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15976
diff changeset
     8
theory Rename imports Extend begin
8256
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
     9
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 16417
diff changeset
    10
definition rename :: "['a => 'b, 'a program] => 'b program" where
8256
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
    11
    "rename h == extend (%(x,u::unit). h x)"
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
    12
63807
5f77017055a3 clarified obscure facts;
wenzelm
parents: 63146
diff changeset
    13
declare image_inv_f_f [simp] image_f_inv_f [simp]
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    14
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    15
declare Extend.intro [simp,intro]
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    16
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    17
lemma good_map_bij [simp,intro]: "bij h ==> good_map (%(x,u). h x)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    18
apply (rule good_mapI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    19
apply (unfold bij_def inj_on_def surj_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    20
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    21
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    22
lemma fst_o_inv_eq_inv: "bij h ==> fst (inv (%(x,u). h x) s) = inv h s"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    23
apply (unfold bij_def split_def, clarify)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    24
apply (subgoal_tac "surj (%p. h (fst p))")
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    25
 prefer 2 apply (simp add: surj_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    26
apply (erule injD)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    27
apply (simp (no_asm_simp) add: surj_f_inv_f)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    28
apply (erule surj_f_inv_f)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    29
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    30
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    31
lemma mem_rename_set_iff: "bij h ==> z \<in> h`A = (inv h z \<in> A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    32
by (force simp add: bij_is_inj bij_is_surj [THEN surj_f_inv_f])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    33
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    34
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    35
lemma extend_set_eq_image [simp]: "extend_set (%(x,u). h x) A = h`A"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    36
by (force simp add: extend_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    37
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    38
lemma Init_rename [simp]: "Init (rename h F) = h`(Init F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    39
by (simp add: rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    40
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    41
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
    42
subsection\<open>inverse properties\<close>
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    43
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    44
lemma extend_set_inv: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    45
     "bij h  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    46
      ==> extend_set (%(x,u::'c). inv h x) = project_set (%(x,u::'c). h x)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    47
apply (unfold bij_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    48
apply (rule ext)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    49
apply (force simp add: extend_set_def project_set_def surj_f_inv_f)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    50
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    51
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    52
(** for "rename" (programs) **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    53
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    54
lemma bij_extend_act_eq_project_act: "bij h  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    55
      ==> extend_act (%(x,u::'c). h x) = project_act (%(x,u::'c). inv h x)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    56
apply (rule ext)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    57
apply (force simp add: extend_act_def project_act_def bij_def surj_f_inv_f)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    58
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    59
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    60
lemma bij_extend_act: "bij h ==> bij (extend_act (%(x,u::'c). h x))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    61
apply (rule bijI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    62
apply (rule Extend.inj_extend_act)
40702
cf26dd7395e4 Replace surj by abbreviation; remove surj_on.
hoelzl
parents: 37936
diff changeset
    63
apply simp
cf26dd7395e4 Replace surj by abbreviation; remove surj_on.
hoelzl
parents: 37936
diff changeset
    64
apply (simp add: bij_extend_act_eq_project_act)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    65
apply (rule surjI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    66
apply (rule Extend.extend_act_inverse)
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45605
diff changeset
    67
apply (blast intro: bij_imp_bij_inv)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    68
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    69
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    70
lemma bij_project_act: "bij h ==> bij (project_act (%(x,u::'c). h x))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    71
apply (frule bij_imp_bij_inv [THEN bij_extend_act])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    72
apply (simp add: bij_extend_act_eq_project_act bij_imp_bij_inv inv_inv_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    73
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    74
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    75
lemma bij_inv_project_act_eq: "bij h ==> inv (project_act (%(x,u::'c). inv h x)) =  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    76
                project_act (%(x,u::'c). h x)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    77
apply (simp (no_asm_simp) add: bij_extend_act_eq_project_act [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    78
apply (rule surj_imp_inv_eq)
40702
cf26dd7395e4 Replace surj by abbreviation; remove surj_on.
hoelzl
parents: 37936
diff changeset
    79
apply (blast intro!: bij_extend_act bij_is_surj)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    80
apply (simp (no_asm_simp) add: Extend.extend_act_inverse)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    81
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    82
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    83
lemma extend_inv: "bij h   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    84
      ==> extend (%(x,u::'c). inv h x) = project (%(x,u::'c). h x) UNIV"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    85
apply (frule bij_imp_bij_inv)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    86
apply (rule ext)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    87
apply (rule program_equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    88
  apply (simp (no_asm_simp) add: extend_set_inv)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    89
 apply (simp add: Extend.project_act_Id Extend.Acts_extend 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    90
          insert_Id_image_Acts bij_extend_act_eq_project_act inv_inv_eq) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    91
apply (simp add: Extend.AllowedActs_extend Extend.AllowedActs_project 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    92
             bij_project_act bij_vimage_eq_inv_image bij_inv_project_act_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    93
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    94
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    95
lemma rename_inv_rename [simp]: "bij h ==> rename (inv h) (rename h F) = F"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    96
by (simp add: rename_def extend_inv Extend.extend_inverse)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    97
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    98
lemma rename_rename_inv [simp]: "bij h ==> rename h (rename (inv h) F) = F"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
    99
apply (frule bij_imp_bij_inv)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   100
apply (erule inv_inv_eq [THEN subst], erule rename_inv_rename)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   101
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   102
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   103
lemma rename_inv_eq: "bij h ==> rename (inv h) = inv (rename h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   104
by (rule inv_equality [symmetric], auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   105
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   106
(** (rename h) is bijective <=> h is bijective **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   107
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   108
lemma bij_extend: "bij h ==> bij (extend (%(x,u::'c). h x))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   109
apply (rule bijI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   110
apply (blast intro: Extend.inj_extend)
15976
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   111
apply (rule_tac f = "extend (% (x,u) . inv h x)" in surjI) 
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   112
apply (subst (1 2) inv_inv_eq [of h, symmetric], assumption+)
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   113
apply (simp add: bij_imp_bij_inv extend_inv [of "inv h"]) 
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   114
apply (simp add: inv_inv_eq)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   115
apply (rule Extend.extend_inverse) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   116
apply (simp add: bij_imp_bij_inv) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   117
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   118
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   119
lemma bij_project: "bij h ==> bij (project (%(x,u::'c). h x) UNIV)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   120
apply (subst extend_inv [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   121
apply (auto simp add: bij_imp_bij_inv bij_extend)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   122
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   123
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   124
lemma inv_project_eq:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   125
     "bij h   
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   126
      ==> inv (project (%(x,u::'c). h x) UNIV) = extend (%(x,u::'c). h x)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   127
apply (rule inj_imp_inv_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   128
apply (erule bij_project [THEN bij_is_inj])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   129
apply (simp (no_asm_simp) add: Extend.extend_inverse)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   130
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   131
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   132
lemma Allowed_rename [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   133
     "bij h ==> Allowed (rename h F) = rename h ` Allowed F"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   134
apply (simp (no_asm_simp) add: rename_def Extend.Allowed_extend)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   135
apply (subst bij_vimage_eq_inv_image)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   136
apply (rule bij_project, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   137
apply (simp (no_asm_simp) add: inv_project_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   138
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   139
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   140
lemma bij_rename: "bij h ==> bij (rename h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   141
apply (simp (no_asm_simp) add: rename_def bij_extend)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   142
done
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 40702
diff changeset
   143
lemmas surj_rename = bij_rename [THEN bij_is_surj]
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   144
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   145
lemma inj_rename_imp_inj: "inj (rename h) ==> inj h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   146
apply (unfold inj_on_def, auto)
15976
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   147
apply (drule_tac x = "mk_program ({x}, {}, {})" in spec)
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   148
apply (drule_tac x = "mk_program ({y}, {}, {})" in spec)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   149
apply (auto simp add: program_equality_iff rename_def extend_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   150
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   151
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   152
lemma surj_rename_imp_surj: "surj (rename h) ==> surj h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   153
apply (unfold surj_def, auto)
15976
44f615d1729b streamlined proof using new subst method
paulson
parents: 15481
diff changeset
   154
apply (drule_tac x = "mk_program ({y}, {}, {})" in spec)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   155
apply (auto simp add: program_equality_iff rename_def extend_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   156
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   157
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   158
lemma bij_rename_imp_bij: "bij (rename h) ==> bij h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   159
apply (unfold bij_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   160
apply (simp (no_asm_simp) add: inj_rename_imp_inj surj_rename_imp_surj)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   161
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   162
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   163
lemma bij_rename_iff [simp]: "bij (rename h) = bij h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   164
by (blast intro: bij_rename bij_rename_imp_bij)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   165
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   166
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   167
subsection\<open>the lattice operations\<close>
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   168
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   169
lemma rename_SKIP [simp]: "bij h ==> rename h SKIP = SKIP"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   170
by (simp add: rename_def Extend.extend_SKIP)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   171
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   172
lemma rename_Join [simp]: 
60773
d09c66a0ea10 more symbols by default, without xsymbols mode;
wenzelm
parents: 58889
diff changeset
   173
     "bij h ==> rename h (F \<squnion> G) = rename h F \<squnion> rename h G"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   174
by (simp add: rename_def Extend.extend_Join)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   175
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   176
lemma rename_JN [simp]:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   177
     "bij h ==> rename h (JOIN I F) = (\<Squnion>i \<in> I. rename h (F i))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   178
by (simp add: rename_def Extend.extend_JN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   179
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   180
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   181
subsection\<open>Strong Safety: co, stable\<close>
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   182
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   183
lemma rename_constrains: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   184
     "bij h ==> (rename h F \<in> (h`A) co (h`B)) = (F \<in> A co B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   185
apply (unfold rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   186
apply (subst extend_set_eq_image [symmetric])+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   187
apply (erule good_map_bij [THEN Extend.intro, THEN Extend.extend_constrains])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   188
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   189
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   190
lemma rename_stable: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   191
     "bij h ==> (rename h F \<in> stable (h`A)) = (F \<in> stable A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   192
apply (simp add: stable_def rename_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   193
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   194
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   195
lemma rename_invariant:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   196
     "bij h ==> (rename h F \<in> invariant (h`A)) = (F \<in> invariant A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   197
apply (simp add: invariant_def rename_stable bij_is_inj [THEN inj_image_subset_iff])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   198
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   199
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   200
lemma rename_increasing:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   201
     "bij h ==> (rename h F \<in> increasing func) = (F \<in> increasing (func o h))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   202
apply (simp add: increasing_def rename_stable [symmetric] bij_image_Collect_eq bij_is_surj [THEN surj_f_inv_f])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   203
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   204
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   205
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   206
subsection\<open>Weak Safety: Co, Stable\<close>
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   207
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   208
lemma reachable_rename_eq: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   209
     "bij h ==> reachable (rename h F) = h ` (reachable F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   210
apply (simp add: rename_def Extend.reachable_extend_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   211
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   212
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   213
lemma rename_Constrains:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   214
     "bij h ==> (rename h F \<in> (h`A) Co (h`B)) = (F \<in> A Co B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   215
by (simp add: Constrains_def reachable_rename_eq rename_constrains
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   216
               bij_is_inj image_Int [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   217
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   218
lemma rename_Stable: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   219
     "bij h ==> (rename h F \<in> Stable (h`A)) = (F \<in> Stable A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   220
by (simp add: Stable_def rename_Constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   221
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   222
lemma rename_Always: "bij h ==> (rename h F \<in> Always (h`A)) = (F \<in> Always A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   223
by (simp add: Always_def rename_Stable bij_is_inj [THEN inj_image_subset_iff])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   224
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   225
lemma rename_Increasing:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   226
     "bij h ==> (rename h F \<in> Increasing func) = (F \<in> Increasing (func o h))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   227
by (simp add: Increasing_def rename_Stable [symmetric] bij_image_Collect_eq 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   228
              bij_is_surj [THEN surj_f_inv_f])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   229
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   230
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   231
subsection\<open>Progress: transient, ensures\<close>
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   232
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   233
lemma rename_transient: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   234
     "bij h ==> (rename h F \<in> transient (h`A)) = (F \<in> transient A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   235
apply (unfold rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   236
apply (subst extend_set_eq_image [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   237
apply (erule good_map_bij [THEN Extend.intro, THEN Extend.extend_transient])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   238
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   239
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   240
lemma rename_ensures: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   241
     "bij h ==> (rename h F \<in> (h`A) ensures (h`B)) = (F \<in> A ensures B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   242
apply (unfold rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   243
apply (subst extend_set_eq_image [symmetric])+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   244
apply (erule good_map_bij [THEN Extend.intro, THEN Extend.extend_ensures])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   245
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   246
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   247
lemma rename_leadsTo: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   248
     "bij h ==> (rename h F \<in> (h`A) leadsTo (h`B)) = (F \<in> A leadsTo B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   249
apply (unfold rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   250
apply (subst extend_set_eq_image [symmetric])+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   251
apply (erule good_map_bij [THEN Extend.intro, THEN Extend.extend_leadsTo])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   252
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   253
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   254
lemma rename_LeadsTo: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   255
     "bij h ==> (rename h F \<in> (h`A) LeadsTo (h`B)) = (F \<in> A LeadsTo B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   256
apply (unfold rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   257
apply (subst extend_set_eq_image [symmetric])+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   258
apply (erule good_map_bij [THEN Extend.intro, THEN Extend.extend_LeadsTo])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   259
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   260
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   261
lemma rename_rename_guarantees_eq: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   262
     "bij h ==> (rename h F \<in> (rename h ` X) guarantees  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   263
                              (rename h ` Y)) =  
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   264
                (F \<in> X guarantees Y)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   265
apply (unfold rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   266
apply (subst good_map_bij [THEN Extend.intro, THEN Extend.extend_guarantees_eq [symmetric]], assumption)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   267
apply (simp (no_asm_simp) add: fst_o_inv_eq_inv o_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   268
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   269
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   270
lemma rename_guarantees_eq_rename_inv:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   271
     "bij h ==> (rename h F \<in> X guarantees Y) =  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   272
                (F \<in> (rename (inv h) ` X) guarantees  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   273
                     (rename (inv h) ` Y))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   274
apply (subst rename_rename_guarantees_eq [symmetric], assumption)
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 60773
diff changeset
   275
apply (simp add: o_def bij_is_surj [THEN surj_f_inv_f] image_comp)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   276
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   277
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   278
lemma rename_preserves:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   279
     "bij h ==> (rename h G \<in> preserves v) = (G \<in> preserves (v o h))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   280
apply (subst good_map_bij [THEN Extend.intro, THEN Extend.extend_preserves [symmetric]], assumption)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   281
apply (simp add: o_def fst_o_inv_eq_inv rename_def bij_is_surj [THEN surj_f_inv_f])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   282
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   283
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   284
lemma ok_rename_iff [simp]: "bij h ==> (rename h F ok rename h G) = (F ok G)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   285
by (simp add: Extend.ok_extend_iff rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   286
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   287
lemma OK_rename_iff [simp]: "bij h ==> OK I (%i. rename h (F i)) = (OK I F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   288
by (simp add: Extend.OK_extend_iff rename_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   289
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   290
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62343
diff changeset
   291
subsection\<open>"image" versions of the rules, for lifting "guarantees" properties\<close>
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   292
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   293
(*All the proofs are similar.  Better would have been to prove one 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   294
  meta-theorem, but how can we handle the polymorphism?  E.g. in 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   295
  rename_constrains the two appearances of "co" have different types!*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   296
lemmas bij_eq_rename = surj_rename [THEN surj_f_inv_f, symmetric]
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   297
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   298
lemma rename_image_constrains:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   299
     "bij h ==> rename h ` (A co B) = (h ` A) co (h`B)" 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   300
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   301
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   302
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   303
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   304
 apply (auto intro!: bij_eq_rename simp add: rename_constrains) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   305
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   306
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   307
lemma rename_image_stable: "bij h ==> rename h ` stable A = stable (h ` A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   308
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   309
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   310
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   311
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   312
 apply (auto intro!: bij_eq_rename simp add: rename_stable)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   313
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   314
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   315
lemma rename_image_increasing:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   316
     "bij h ==> rename h ` increasing func = increasing (func o inv h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   317
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   318
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   319
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   320
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   321
 apply (auto intro!: bij_eq_rename simp add: rename_increasing o_def bij_is_inj) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   322
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   323
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   324
lemma rename_image_invariant:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   325
     "bij h ==> rename h ` invariant A = invariant (h ` A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   326
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   327
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   328
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   329
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   330
 apply (auto intro!: bij_eq_rename simp add: rename_invariant) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   331
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   332
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   333
lemma rename_image_Constrains:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   334
     "bij h ==> rename h ` (A Co B) = (h ` A) Co (h`B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   335
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   336
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   337
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   338
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   339
 apply (auto intro!: bij_eq_rename simp add: rename_Constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   340
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   341
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   342
lemma rename_image_preserves:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   343
     "bij h ==> rename h ` preserves v = preserves (v o inv h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   344
by (simp add: o_def rename_image_stable preserves_def bij_image_INT 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   345
              bij_image_Collect_eq)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   346
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   347
lemma rename_image_Stable:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   348
     "bij h ==> rename h ` Stable A = Stable (h ` A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   349
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   350
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   351
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   352
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   353
 apply (auto intro!: bij_eq_rename simp add: rename_Stable) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   354
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   355
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   356
lemma rename_image_Increasing:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   357
     "bij h ==> rename h ` Increasing func = Increasing (func o inv h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   358
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   359
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   360
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   361
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   362
 apply (auto intro!: bij_eq_rename simp add: rename_Increasing o_def bij_is_inj)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   363
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   364
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   365
lemma rename_image_Always: "bij h ==> rename h ` Always A = Always (h ` A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   366
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   367
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   368
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   369
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   370
 apply (auto intro!: bij_eq_rename simp add: rename_Always)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   371
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   372
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   373
lemma rename_image_leadsTo:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   374
     "bij h ==> rename h ` (A leadsTo B) = (h ` A) leadsTo (h`B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   375
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   376
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   377
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   378
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   379
 apply (auto intro!: bij_eq_rename simp add: rename_leadsTo) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   380
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   381
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   382
lemma rename_image_LeadsTo:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   383
     "bij h ==> rename h ` (A LeadsTo B) = (h ` A) LeadsTo (h`B)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   384
apply auto 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   385
 defer 1
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   386
 apply (rename_tac F) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   387
 apply (subgoal_tac "\<exists>G. F = rename h G") 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   388
 apply (auto intro!: bij_eq_rename simp add: rename_LeadsTo) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   389
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 9403
diff changeset
   390
8256
6ba8fa2b0638 Rename: theory for applying a bijection over states to a UNITY program
paulson
parents:
diff changeset
   391
end