Nominal/Ex/Pi.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 15 Jul 2012 13:03:47 +0100
changeset 3192 14c7d7e29c44
parent 3191 0440bc1a2438
child 3197 25d11b449e92
permissions -rw-r--r--
added a simproc for alpha-equivalence to the simplifier
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3097
b27e94db1b8a included Pi theory in tests
Christian Urban <urbanc@in.tum.de>
parents: 3096
diff changeset
     1
(* Theory be Kirstin Peters *)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
3097
b27e94db1b8a included Pi theory in tests
Christian Urban <urbanc@in.tum.de>
parents: 3096
diff changeset
     3
theory Pi
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
  imports "../Nominal2"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
begin
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
atom_decl name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
subsection {* Capture-Avoiding Substitution of Names *}
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
definition
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  subst_name :: "name \<Rightarrow> name \<Rightarrow> name \<Rightarrow> name" ("_[_:::=_]" [110, 110, 110] 110)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
where
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  "a[b:::=c] \<equiv> if (a = b) then c else a"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
declare subst_name_def[simp]
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
lemma subst_name_mix_eqvt[eqvt]:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  fixes p :: perm
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  and   a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  and   b :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  and   c :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  shows "p \<bullet> (a[b:::=c]) = (p \<bullet> a)[(p \<bullet> b):::=(p \<bullet> c)]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
proof -
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  show ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
    by(auto)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
nominal_primrec
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  subst_name_list :: "name \<Rightarrow> (name \<times> name) list \<Rightarrow> name"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
where
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  "subst_name_list a [] = a"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
| "subst_name_list a ((b, c)#xs) = (if (a = b) then c else (subst_name_list a xs))"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  apply(auto)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  apply(subgoal_tac "\<And>p x r. subst_name_list_graph x r \<Longrightarrow> subst_name_list_graph (p \<bullet> x) (p \<bullet> r)")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  unfolding eqvt_def
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
    38
  apply(simp only: permute_fun_def)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  apply(rule allI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  apply(rule ext)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  apply(rule ext)
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
    42
  apply(simp only: permute_bool_def)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  apply(rule iffI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  apply(drule_tac x="p" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  apply(drule_tac x="- p \<bullet> x" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  apply(drule_tac x="- p \<bullet> xa" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  apply(drule_tac x="-p" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  apply(drule_tac x="x" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  apply(drule_tac x="xa" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  apply(erule subst_name_list_graph.induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  apply(perm_simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  apply(rule subst_name_list_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  apply(perm_simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  apply(rule subst_name_list_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  apply(rule_tac y="b" in list.exhaust)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  by(auto)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
termination (eqvt)
3097
b27e94db1b8a included Pi theory in tests
Christian Urban <urbanc@in.tum.de>
parents: 3096
diff changeset
    62
  by (lexicographic_order)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
section {* The Synchronous Pi-Calculus *}
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
subsection {* Syntax: Synchronous, Monadic Pi-Calculus with n-ary, Mixed Choice *}
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
nominal_datatype
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
      guardedTerm_mix = Output name name piMix                     ("_!<_>\<onesuperior>._" [120, 120, 110] 110)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
                      | Input name b::name P::piMix  binds b in P  ("_?<_>\<onesuperior>._" [120, 120, 110] 110)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
                      | Tau piMix                                  ("<\<tau>\<onesuperior>>._" [110] 110)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  and sumList_mix     = SumNil                                     ("\<zero>\<onesuperior>")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
                      | AddSummand guardedTerm_mix sumList_mix     (infixr "\<oplus>\<onesuperior>" 65)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  and piMix           = Res a::name P::piMix         binds a in P  ("<\<nu>_>\<onesuperior>_" [100, 100] 100)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
                      | Par piMix piMix                            (infixr "\<parallel>\<onesuperior>" 85)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
                      | Match name name piMix                      ("[_\<frown>\<onesuperior>_]_" [120, 120, 110] 110)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
                      | Sum sumList_mix                            ("\<oplus>\<onesuperior>{_}" 90)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
                      | Rep name b::name P::piMix    binds b in P  ("\<infinity>_?<_>\<onesuperior>._" [120, 120, 110] 110)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
                      | Succ                                       ("succ\<onesuperior>")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
lemmas piMix_strong_induct  = guardedTerm_mix_sumList_mix_piMix.strong_induct
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
lemmas piMix_fresh          = guardedTerm_mix_sumList_mix_piMix.fresh
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
lemmas piMix_eq_iff         = guardedTerm_mix_sumList_mix_piMix.eq_iff
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
lemmas piMix_distinct       = guardedTerm_mix_sumList_mix_piMix.distinct
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
lemmas piMix_size           = guardedTerm_mix_sumList_mix_piMix.size
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
subsection {* Alpha-Conversion Lemmata *}
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
lemma alphaRes_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
  fixes a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
  and   P :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  and   z :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  assumes "atom z \<sharp> P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  shows "<\<nu>a>\<onesuperior>P = <\<nu>z>\<onesuperior>((atom a \<rightleftharpoons> atom z) \<bullet> P)"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
proof(cases "a = z")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  assume "a = z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
    by(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
next
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  assume "a \<noteq> z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
    using assms
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   106
    by (simp add: flip_def piMix_eq_iff Abs1_eq_iff fresh_permute_left)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
lemma alphaInput_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  fixes a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  and   b :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  and   P :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  and   z :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  assumes "atom z \<sharp> P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  shows "a?<b>\<onesuperior>.P = a?<z>\<onesuperior>.((atom b \<rightleftharpoons> atom z) \<bullet> P)"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
proof(cases "b = z")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  assume "b = z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
    by(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
next
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  assume "b \<noteq> z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
    using assms
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   126
    by(simp add: flip_def piMix_eq_iff Abs1_eq_iff fresh_permute_left)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
lemma alphaRep_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  fixes a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  and   b :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  and   P :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  and   z :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  assumes "atom z \<sharp> P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
  shows "\<infinity>a?<b>\<onesuperior>.P = \<infinity>a?<z>\<onesuperior>.((atom b \<rightleftharpoons> atom z) \<bullet> P)"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
proof(cases "b = z")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  assume "b = z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
    by(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
next
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  assume "b \<noteq> z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
    using assms
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   146
    by(simp add: flip_def piMix_eq_iff Abs1_eq_iff fresh_permute_left)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
subsection {* Capture-Avoiding Substitution of Names *}
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
lemma testl:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  assumes a: "\<exists>y. f = Inl y"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  shows "(p \<bullet> (Sum_Type.Projl f)) = Sum_Type.Projl (p \<bullet> f)"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
using a by auto
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
lemma testrr:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  assumes a: "\<exists>y. f = Inr (Inr y)"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  shows "(p \<bullet> (Sum_Type.Projr (Sum_Type.Projr f))) = Sum_Type.Projr (Sum_Type.Projr (p \<bullet> f))"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
using a by auto
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
lemma testlr:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  assumes a: "\<exists>y. f = Inr (Inl y)"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
  shows "(p \<bullet> (Sum_Type.Projl (Sum_Type.Projr f))) = Sum_Type.Projl (Sum_Type.Projr (p \<bullet> f))"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
using a by auto
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
nominal_primrec (default "sum_case (\<lambda>x. Inl undefined) (sum_case (\<lambda>x. Inr (Inl undefined)) (\<lambda>x. Inr (Inr undefined)))")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  subsGuard_mix :: "guardedTerm_mix \<Rightarrow> name \<Rightarrow> name \<Rightarrow> guardedTerm_mix"  ("_[_::=\<onesuperior>\<onesuperior>_]" [100, 100, 100] 100) and
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  subsList_mix  :: "sumList_mix \<Rightarrow> name \<Rightarrow> name \<Rightarrow> sumList_mix"          ("_[_::=\<onesuperior>\<twosuperior>_]" [100, 100, 100] 100) and
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  subs_mix      :: "piMix \<Rightarrow> name \<Rightarrow> name \<Rightarrow> piMix"                      ("_[_::=\<onesuperior>_]" [100, 100, 100] 100)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
where
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  "(a!<b>\<onesuperior>.P)[x::=\<onesuperior>\<onesuperior>y] = (a[x:::=y])!<(b[x:::=y])>\<onesuperior>.(P[x::=\<onesuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
| "\<lbrakk>atom b \<sharp> (x, y)\<rbrakk> \<Longrightarrow> (a?<b>\<onesuperior>.P)[x::=\<onesuperior>\<onesuperior>y] = (a[x:::=y])?<b>\<onesuperior>.(P[x::=\<onesuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
| "(<\<tau>\<onesuperior>>.P)[x::=\<onesuperior>\<onesuperior>y] = <\<tau>\<onesuperior>>.(P[x::=\<onesuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
| "(\<zero>\<onesuperior>)[x::=\<onesuperior>\<twosuperior>y] = \<zero>\<onesuperior>"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
| "(g \<oplus>\<onesuperior> xg)[x::=\<onesuperior>\<twosuperior>y] = (g[x::=\<onesuperior>\<onesuperior>y]) \<oplus>\<onesuperior> (xg[x::=\<onesuperior>\<twosuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
| "\<lbrakk>atom a \<sharp> (x, y)\<rbrakk> \<Longrightarrow> (<\<nu>a>\<onesuperior>P)[x::=\<onesuperior>y] = <\<nu>a>\<onesuperior>(P[x::=\<onesuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
| "(P \<parallel>\<onesuperior> Q)[x::=\<onesuperior>y] = (P[x::=\<onesuperior>y]) \<parallel>\<onesuperior> (Q[x::=\<onesuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
| "([a\<frown>\<onesuperior>b]P)[x::=\<onesuperior>y] = ([(a[x:::=y])\<frown>\<onesuperior>(b[x:::=y])](P[x::=\<onesuperior>y]))"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
| "(\<oplus>\<onesuperior>{xg})[x::=\<onesuperior>y] = \<oplus>\<onesuperior>{(xg[x::=\<onesuperior>\<twosuperior>y])}"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
| "\<lbrakk>atom b \<sharp> (x, y)\<rbrakk> \<Longrightarrow> (\<infinity>a?<b>\<onesuperior>.P)[x::=\<onesuperior>y] = \<infinity>(a[x:::=y])?<b>\<onesuperior>.(P[x::=\<onesuperior>y])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
| "(succ\<onesuperior>)[x::=\<onesuperior>y] = succ\<onesuperior>"
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
   182
  using [[simproc del: alpha_lst]]
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
  apply(auto simp add: piMix_distinct piMix_eq_iff)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  apply(subgoal_tac "\<And>p x r. subsGuard_mix_subsList_mix_subs_mix_graph x r \<Longrightarrow> subsGuard_mix_subsList_mix_subs_mix_graph (p \<bullet> x) (p \<bullet> r)")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  unfolding eqvt_def
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  apply(rule allI)
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   187
  apply(simp only: permute_fun_def)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  apply(rule ext)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
  apply(rule ext)
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   190
  apply(simp only: permute_bool_def)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  apply(rule iffI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  apply(drule_tac x="p" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  apply(drule_tac x="- p \<bullet> x" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  apply(drule_tac x="- p \<bullet> xa" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  apply(drule_tac x="-p" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
  apply(drule_tac x="x" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  apply(drule_tac x="xa" in meta_spec)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  --"Equivariance"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  apply(simp only: atom_eqvt[symmetric] Pair_eqvt[symmetric] fresh_eqvt[symmetric] permute_bool_def)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  apply(simp (no_asm_use) only: eqvts)  
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
  apply(subst testl)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  apply(subst testlr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  apply(rotate_tac 2)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
  apply(perm_simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
  apply(simp only: atom_eqvt[symmetric] Pair_eqvt[symmetric] fresh_eqvt[symmetric] permute_bool_def)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  apply(rotate_tac 2)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  apply(perm_simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
  apply(subst testlr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  apply(simp (no_asm_use) only: eqvts)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  apply(subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
  apply(erule subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  apply(blast)+
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  apply(simp only: atom_eqvt[symmetric] Pair_eqvt[symmetric] fresh_eqvt[symmetric] permute_bool_def)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  apply(perm_simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  apply(rule subsGuard_mix_subsList_mix_subs_mix_graph.intros)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  --"Covered all cases"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  apply(case_tac x)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
  apply(case_tac a)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
  apply (rule_tac y="aa" and c="(b, c)" in guardedTerm_mix_sumList_mix_piMix.strong_exhaust(1))
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  apply(auto simp add: fresh_star_def)[1]
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
  apply(case_tac b)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  apply(case_tac a)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  apply (rule_tac ya="aa" in guardedTerm_mix_sumList_mix_piMix.strong_exhaust(2))
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
  apply(case_tac ba)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
  apply (rule_tac yb="a" and c="(bb,c)" in guardedTerm_mix_sumList_mix_piMix.strong_exhaust(3))
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
   298
  using [[simproc del: alpha_lst]]
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
  apply(auto simp add: fresh_star_def)[1]
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
  apply(blast)
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
   303
  using [[simproc del: alpha_lst]]
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  apply(auto simp add: fresh_star_def)[1]
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
  apply(simp)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  apply(blast)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  --"compatibility"
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   309
  apply (simp only: meta_eq_to_obj_eq[OF subs_mix_def, symmetric, unfolded fun_eq_iff])
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
  apply (subgoal_tac "eqvt_at (\<lambda>(a, b, c). subs_mix a b c) (P, xa, ya)")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
  apply (thin_tac "eqvt_at subsGuard_mix_subsList_mix_subs_mix_sumC (Inr (Inr (P, xa, ya)))")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
  apply (thin_tac "eqvt_at subsGuard_mix_subsList_mix_subs_mix_sumC (Inr (Inr (Pa, xa, ya)))")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  prefer 2
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   314
  apply (simp only: eqvt_at_def subs_mix_def)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
  apply rule
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   316
  apply(simp (no_asm))
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
  apply (subst testrr)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
  apply (simp add: subsGuard_mix_subsList_mix_subs_mix_sumC_def)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
  apply (simp add: THE_default_def)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
apply (case_tac "Ex1 (subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (P, xa, ya))))")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
apply simp_all[2]
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
apply auto[1]
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
apply (erule_tac x="x" in allE)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
apply simp
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
apply (thin_tac "\<forall>p\<Colon>perm.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
           p \<bullet> The (subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (P, xa, ya)))) =
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
           (if \<exists>!x\<Colon>guardedTerm_mix + sumList_mix + piMix.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
                  subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \<bullet> P, p \<bullet> xa, p \<bullet> ya))) x
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
            then THE x\<Colon>guardedTerm_mix + sumList_mix + piMix.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
                    subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \<bullet> P, p \<bullet> xa, p \<bullet> ya))) x
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
            else Inr (Inr undefined))")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
apply (thin_tac "\<forall>p\<Colon>perm.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
           p \<bullet> (if \<exists>!x\<Colon>guardedTerm_mix + sumList_mix + piMix.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
                      subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (Pa, xa, ya))) x
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
                then THE x\<Colon>guardedTerm_mix + sumList_mix + piMix.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
                        subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (Pa, xa, ya))) x
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
                else Inr (Inr undefined)) =
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
           (if \<exists>!x\<Colon>guardedTerm_mix + sumList_mix + piMix.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
                  subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \<bullet> Pa, p \<bullet> xa, p \<bullet> ya))) x
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
            then THE x\<Colon>guardedTerm_mix + sumList_mix + piMix.
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
                    subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \<bullet> Pa, p \<bullet> xa, p \<bullet> ya))) x
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
            else Inr (Inr undefined))")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
apply (thin_tac "atom b \<sharp> (xa, ya)")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
apply (thin_tac "atom ba \<sharp> (xa, ya)")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
apply (thin_tac "[[atom b]]lst. P = [[atom ba]]lst. Pa")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
apply(cases rule: subsGuard_mix_subsList_mix_subs_mix_graph.cases)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
apply (metis Inr_not_Inl)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
apply (metis Inr_not_Inl)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
apply (metis Inr_not_Inl)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
apply (metis Inr_inject Inr_not_Inl)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
apply (metis Inr_inject Inr_not_Inl)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
apply (rule_tac x="<\<nu>a>\<onesuperior>Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
                            (Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
                              (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y)))))" in exI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
apply clarify
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
apply (rule the1_equality)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
apply blast apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
apply (rule_tac x="Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
                       (Sum_Type.Projr (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y))))) \<parallel>\<onesuperior>
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
                      Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
                       (Sum_Type.Projr (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Q, xb, y)))))" in exI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
apply clarify
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
apply (rule the1_equality)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
apply blast apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
apply (rule_tac x="[(a[xb:::=y])\<frown>\<onesuperior>(bb[xb:::=y])]Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
                                                    (Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
(subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y)))))" in exI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
apply clarify
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
apply (rule the1_equality)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
apply blast apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
apply (rule_tac x="\<oplus>\<onesuperior>{Sum_Type.Projl
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
                          (Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
                            (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inl (xg, xb, y)))))}" in exI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
apply clarify
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
apply (rule the1_equality)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
apply blast apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
apply (rule_tac x="\<infinity>(a[xb:::=y])?<bb>\<onesuperior>.Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
                                           (Sum_Type.Projr
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
                                             (subsGuard_mix_subsList_mix_subs_mix_sum
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
                                               (Inr (Inr (Pb, xb, y)))))" in exI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
apply clarify
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
apply (rule the1_equality)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
apply blast apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
apply (rule_tac x="succ\<onesuperior>" in exI)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
apply clarify
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
apply (rule the1_equality)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
apply blast apply assumption
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
apply simp
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
(* Here the only real goal compatibility is left *)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
  apply (erule Abs_lst1_fcb)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
  apply (simp_all add: Abs_fresh_iff fresh_fun_eqvt_app)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
  apply (subgoal_tac "atom ba \<sharp> (\<lambda>(a, x, y). subs_mix a x y) (P, xa, ya)")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
  apply simp
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
  apply (erule fresh_eqvt_at)
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   396
  apply(simp add: finite_supp)
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   397
  apply(simp)
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   398
  apply(simp add: eqvt_at_def)
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   399
  apply(drule_tac x="(b \<leftrightarrow> ba)" in spec)
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3097
diff changeset
   400
  apply(simp)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
  done
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
termination (eqvt)
3097
b27e94db1b8a included Pi theory in tests
Christian Urban <urbanc@in.tum.de>
parents: 3096
diff changeset
   404
  by (lexicographic_order)
3096
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
lemma forget_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
  fixes g  :: guardedTerm_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
  and   xg :: sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
  and   P  :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
  and   x  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
  and   y  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
  shows "atom x \<sharp> g \<longrightarrow> g[x::=\<onesuperior>\<onesuperior>y] = g"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
  and   "atom x \<sharp> xg \<longrightarrow> xg[x::=\<onesuperior>\<twosuperior>y] = xg"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
  and   "atom x \<sharp> P \<longrightarrow> P[x::=\<onesuperior>y] = P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
proof -
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
  show  "atom x \<sharp> g \<longrightarrow> g[x::=\<onesuperior>\<onesuperior>y] = g"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
  and   "atom x \<sharp> xg \<longrightarrow> xg[x::=\<onesuperior>\<twosuperior>y] = xg"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
  and   "atom x \<sharp> P \<longrightarrow> P[x::=\<onesuperior>y] = P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
    using assms
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
    apply(nominal_induct g and xg and P avoiding: x y rule: piMix_strong_induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
    by(auto simp add: piMix_eq_iff piMix_fresh fresh_at_base)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
lemma fresh_fact_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
  fixes g  :: guardedTerm_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
  and   xg :: sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  and   P  :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
  and   x  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
  and   y  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
  and   z  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
  assumes "atom z \<sharp> y"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
  shows "(z = x \<or> atom z \<sharp> g) \<longrightarrow> atom z \<sharp> g[x::=\<onesuperior>\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
  and   "(z = x \<or> atom z \<sharp> xg) \<longrightarrow> atom z \<sharp> xg[x::=\<onesuperior>\<twosuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
  and   "(z = x \<or> atom z \<sharp> P) \<longrightarrow> atom z \<sharp> P[x::=\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
proof -
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  show  "(z = x \<or> atom z \<sharp> g) \<longrightarrow> atom z \<sharp> g[x::=\<onesuperior>\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
  and   "(z = x \<or> atom z \<sharp> xg) \<longrightarrow> atom z \<sharp> xg[x::=\<onesuperior>\<twosuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
  and   "(z = x \<or> atom z \<sharp> P) \<longrightarrow> atom z \<sharp> P[x::=\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
    using assms
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
    apply(nominal_induct g and xg and P avoiding: x y z rule: piMix_strong_induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
    by(auto simp add: piMix_fresh fresh_at_base)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
lemma substitution_lemma_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
  fixes g  :: guardedTerm_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
  and   xg :: sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
  and   P  :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
  and   s  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
  and   u  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
  and   x  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
  and   y  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
  assumes "x \<noteq> y"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
  and     "atom x \<sharp> u"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
  shows "g[x::=\<onesuperior>\<onesuperior>s][y::=\<onesuperior>\<onesuperior>u] = g[y::=\<onesuperior>\<onesuperior>u][x::=\<onesuperior>\<onesuperior>s[y:::=u]]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
  and   "xg[x::=\<onesuperior>\<twosuperior>s][y::=\<onesuperior>\<twosuperior>u] = xg[y::=\<onesuperior>\<twosuperior>u][x::=\<onesuperior>\<twosuperior>s[y:::=u]]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
  and   "P[x::=\<onesuperior>s][y::=\<onesuperior>u] = P[y::=\<onesuperior>u][x::=\<onesuperior>s[y:::=u]]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
proof -
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
  show  "g[x::=\<onesuperior>\<onesuperior>s][y::=\<onesuperior>\<onesuperior>u] = g[y::=\<onesuperior>\<onesuperior>u][x::=\<onesuperior>\<onesuperior>s[y:::=u]]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  and   "xg[x::=\<onesuperior>\<twosuperior>s][y::=\<onesuperior>\<twosuperior>u] = xg[y::=\<onesuperior>\<twosuperior>u][x::=\<onesuperior>\<twosuperior>s[y:::=u]]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
  and   "P[x::=\<onesuperior>s][y::=\<onesuperior>u] = P[y::=\<onesuperior>u][x::=\<onesuperior>s[y:::=u]]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
    using assms
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
    apply(nominal_induct g and xg and P avoiding: x y s u rule: piMix_strong_induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
    apply(simp_all add: fresh_fact_mix forget_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
    by(auto simp add: fresh_at_base)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
lemma perm_eq_subst_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  fixes g  :: guardedTerm_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
  and   xg :: sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
  and   P  :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
  and   x  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
  and   y  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
  shows "atom y \<sharp> g \<longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> g = g[x::=\<onesuperior>\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
  and   "atom y \<sharp> xg \<longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> xg = xg[x::=\<onesuperior>\<twosuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
  and   "atom y \<sharp> P \<longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> P = P[x::=\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
proof -
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
  show  "atom y \<sharp> g \<longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> g = g[x::=\<onesuperior>\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
  and   "atom y \<sharp> xg \<longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> xg = xg[x::=\<onesuperior>\<twosuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
  and   "atom y \<sharp> P \<longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> P = P[x::=\<onesuperior>y]"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
    apply(nominal_induct g and xg and P avoiding: x y rule: piMix_strong_induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
    by(auto simp add: piMix_fresh fresh_at_base)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
lemma subst_id_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
  fixes g  :: guardedTerm_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
  and   xg :: sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
  and   P  :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
  and   x  :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
  shows "g[x::=\<onesuperior>\<onesuperior>x] = g" and "xg[x::=\<onesuperior>\<twosuperior>x] = xg" and "P[x::=\<onesuperior>x] = P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
proof -
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
  show  "g[x::=\<onesuperior>\<onesuperior>x] = g" and "xg[x::=\<onesuperior>\<twosuperior>x] = xg" and "P[x::=\<onesuperior>x] = P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
    apply(nominal_induct g and xg and P avoiding: x rule: piMix_strong_induct)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
    by(auto)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
lemma alphaRes_subst_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
  fixes a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
  and   P :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
  and   z :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
  assumes "atom z \<sharp> P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
  shows "<\<nu>a>\<onesuperior>P = <\<nu>z>\<onesuperior>(P[a::=\<onesuperior>z])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
proof(cases "a = z")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
  assume "a = z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
    by(simp add: subst_id_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
next
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
  assume "a \<noteq> z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
    using assms
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
    by(simp add: alphaRes_mix perm_eq_subst_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
lemma alphaInput_subst_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  fixes a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
  and   b :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
  and   P :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
  and   z :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
  assumes "atom z \<sharp> P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
  shows "a?<b>\<onesuperior>.P = a?<z>\<onesuperior>.(P[b::=\<onesuperior>z])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
proof(cases "b = z")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
  assume "b = z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
    by(simp add: subst_id_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
next
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
  assume "b \<noteq> z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
    using assms
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
    by(simp add: alphaInput_mix perm_eq_subst_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
lemma alphaRep_subst_mix:
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
  fixes a :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
  and   b :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
  and   P :: piMix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
  and   z :: name
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
  assumes "atom z \<sharp> P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  shows "\<infinity>a?<b>\<onesuperior>.P = \<infinity>a?<z>\<onesuperior>.(P[b::=\<onesuperior>z])"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
proof(cases "b = z")
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
  assume "b = z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
    by(simp add: subst_id_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
next
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  assume "b \<noteq> z"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
  thus ?thesis
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
    using assms
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
    by(simp add: alphaRep_mix perm_eq_subst_mix)
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
qed
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
inductive
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
  fresh_list_guard_mix :: "name list \<Rightarrow> guardedTerm_mix \<Rightarrow> bool"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
where
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
  "fresh_list_guard_mix [] g"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
| "\<lbrakk>atom n \<sharp> g; fresh_list_guard_mix xn g\<rbrakk> \<Longrightarrow> fresh_list_guard_mix (n#xn) g"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
equivariance fresh_list_guard_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
nominal_inductive fresh_list_guard_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
  done
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
inductive
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
  fresh_list_sumList_mix :: "name list \<Rightarrow> sumList_mix \<Rightarrow> bool"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
where
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
  "fresh_list_sumList_mix [] xg"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
| "\<lbrakk>atom n \<sharp> xg; fresh_list_sumList_mix xn xg\<rbrakk> \<Longrightarrow> fresh_list_sumList_mix (n#xn) xg"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
equivariance fresh_list_sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
nominal_inductive fresh_list_sumList_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  done
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
inductive
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
  fresh_list_mix :: "name list \<Rightarrow> piMix \<Rightarrow> bool"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
where
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
  "fresh_list_mix [] P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
| "\<lbrakk>atom n \<sharp> P; fresh_list_mix xn P\<rbrakk> \<Longrightarrow> fresh_list_mix (n#xn) P"
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
equivariance fresh_list_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
nominal_inductive fresh_list_mix
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
  done
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
18f20d75b463 added file by Kirstin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
end