Nominal/Ex/CPS/CPS1_Plotkin.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Sun, 24 Jul 2011 07:54:54 +0200
changeset 2984 1b39ba5db2c1
parent 2933 3be019a86117
child 2998 f0fab367453a
permissions -rw-r--r--
update to 'termination (eqvt)'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
header {* CPS conversion *}
2895
65efa1c7563c Theory name changes for JEdit
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
     2
theory CPS1_Plotkin
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
imports Lt
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
begin
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
2933
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
     6
lemma Abs_lst_fcb2:
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
     7
  fixes as bs :: "atom list"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
     8
    and x y :: "'b :: fs"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
     9
    and c::"'c::fs"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    10
  assumes eq: "[as]lst. x = [bs]lst. y"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    11
  and fcb1: "(set as) \<sharp>* c \<Longrightarrow> (set as) \<sharp>* f as x c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    12
  and fresh1: "set as \<sharp>* c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    13
  and fresh2: "set bs \<sharp>* c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    14
  and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    15
  and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    16
  shows "f as x c = f bs y c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    17
proof -
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    18
  have "supp (as, x, c) supports (f as x c)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    19
    unfolding  supports_def fresh_def[symmetric]
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    20
    by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    21
  then have fin1: "finite (supp (f as x c))"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    22
    by (auto intro: supports_finite simp add: finite_supp)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    23
  have "supp (bs, y, c) supports (f bs y c)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    24
    unfolding  supports_def fresh_def[symmetric]
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    25
    by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    26
  then have fin2: "finite (supp (f bs y c))"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    27
    by (auto intro: supports_finite simp add: finite_supp)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    28
  obtain q::"perm" where 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    29
    fr1: "(q \<bullet> (set as)) \<sharp>* (x, c, f as x c, f bs y c)" and 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    30
    fr2: "supp q \<sharp>* Abs_lst as x" and 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    31
    inc: "supp q \<subseteq> (set as) \<union> q \<bullet> (set as)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    32
    using at_set_avoiding3[where xs="set as" and c="(x, c, f as x c, f bs y c)" and x="[as]lst. x"]  
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    33
      fin1 fin2
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    34
    by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    35
  have "Abs_lst (q \<bullet> as) (q \<bullet> x) = q \<bullet> Abs_lst as x" by simp
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    36
  also have "\<dots> = Abs_lst as x"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    37
    by (simp only: fr2 perm_supp_eq)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    38
  finally have "Abs_lst (q \<bullet> as) (q \<bullet> x) = Abs_lst bs y" using eq by simp
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    39
  then obtain r::perm where 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    40
    qq1: "q \<bullet> x = r \<bullet> y" and 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    41
    qq2: "q \<bullet> as = r \<bullet> bs" and 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    42
    qq3: "supp r \<subseteq> (q \<bullet> (set as)) \<union> set bs"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    43
    apply(drule_tac sym)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    44
    apply(simp only: Abs_eq_iff2 alphas)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    45
    apply(erule exE)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    46
    apply(erule conjE)+
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    47
    apply(drule_tac x="p" in meta_spec)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    48
    apply(simp add: set_eqvt)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    49
    apply(blast)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    50
    done
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    51
  have "(set as) \<sharp>* f as x c" 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    52
    apply(rule fcb1)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    53
    apply(rule fresh1)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    54
    done
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    55
  then have "q \<bullet> ((set as) \<sharp>* f as x c)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    56
    by (simp add: permute_bool_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    57
  then have "set (q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    58
    apply(simp add: fresh_star_eqvt set_eqvt)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    59
    apply(subst (asm) perm1)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    60
    using inc fresh1 fr1
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    61
    apply(auto simp add: fresh_star_def fresh_Pair)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    62
    done
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    63
  then have "set (r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    64
  then have "r \<bullet> ((set bs) \<sharp>* f bs y c)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    65
    apply(simp add: fresh_star_eqvt set_eqvt)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    66
    apply(subst (asm) perm2[symmetric])
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    67
    using qq3 fresh2 fr1
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    68
    apply(auto simp add: set_eqvt fresh_star_def fresh_Pair)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    69
    done
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    70
  then have fcb2: "(set bs) \<sharp>* f bs y c" by (simp add: permute_bool_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    71
  have "f as x c = q \<bullet> (f as x c)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    72
    apply(rule perm_supp_eq[symmetric])
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    73
    using inc fcb1[OF fresh1] fr1 by (auto simp add: fresh_star_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    74
  also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" 
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    75
    apply(rule perm1)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    76
    using inc fresh1 fr1 by (auto simp add: fresh_star_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    77
  also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    78
  also have "\<dots> = r \<bullet> (f bs y c)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    79
    apply(rule perm2[symmetric])
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    80
    using qq3 fresh2 fr1 by (auto simp add: fresh_star_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    81
  also have "... = f bs y c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    82
    apply(rule perm_supp_eq)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    83
    using qq3 fr1 fcb2 by (auto simp add: fresh_star_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    84
  finally show ?thesis by simp
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    85
qed
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    86
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    87
lemma Abs_lst1_fcb2:
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    88
  fixes a b :: "atom"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    89
    and x y :: "'b :: fs"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    90
    and c::"'c :: fs"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    91
  assumes e: "(Abs_lst [a] x) = (Abs_lst [b] y)"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    92
  and fcb1: "a \<sharp> c \<Longrightarrow> a \<sharp> f a x c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    93
  and fresh: "{a, b} \<sharp>* c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    94
  and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    95
  and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    96
  shows "f a x c = f b y c"
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    97
using e
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    98
apply(drule_tac Abs_lst_fcb2[where c="c" and f="\<lambda>(as::atom list) . f (hd as)"])
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    99
apply(simp_all)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
   100
using fcb1 fresh perm1 perm2
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
   101
apply(simp_all add: fresh_star_def)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
   102
done
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
   103
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
nominal_primrec
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
  CPS :: "lt \<Rightarrow> lt" ("_*" [250] 250)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
where
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
  "atom k \<sharp> x \<Longrightarrow> (x~)* = (Abs k ((k~) $ (x~)))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
| "atom k \<sharp> (x, M) \<Longrightarrow> (Abs x M)* = Abs k (k~ $ Abs x (M*))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
| "atom k \<sharp> (M, N) \<Longrightarrow> atom m \<sharp> (N, k) \<Longrightarrow> atom n \<sharp> (k, m) \<Longrightarrow>
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
    (M $ N)* = Abs k (M* $ Abs m (N* $ Abs n (m~ $ n~ $ k~)))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
unfolding eqvt_def CPS_graph_def
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
apply (rule, perm_simp, rule, rule)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
apply (simp_all add: fresh_Pair_elim)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
apply (rule_tac y="x" in lt.exhaust)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
apply (auto)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
apply (rule_tac x="name" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
apply (simp_all add: fresh_at_base)[3]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
apply (rule_tac x="(name, lt)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
apply (simp add: fresh_Pair_elim fresh_at_base)[2]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
apply (rule_tac x="(lt1, lt2)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
apply (rule_tac x="(lt2, a)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
apply (rule_tac x="(a, aa)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
apply (simp add: fresh_Pair_elim fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
apply (simp add: Abs1_eq_iff lt.fresh fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
--"-"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
apply(rule_tac s="[[atom ka]]lst. ka~ $ Abs x (CPS_sumC M)" in trans)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
apply (case_tac "k = ka")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
apply simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
apply(simp (no_asm) add: Abs1_eq_iff del:eqvts)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
apply (simp del: eqvts add: lt.fresh fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
apply (simp only: lt.perm_simps(1) lt.perm_simps(3) flip_def[symmetric] lt.eq_iff(3))
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
apply (subst  flip_at_base_simps(2))
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
apply simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
apply (intro conjI refl)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
apply (rule flip_fresh_fresh[symmetric])
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
apply (simp_all add: lt.fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
apply (metis fresh_eqvt_at lt.fsupp)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
apply (case_tac "ka = x")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
apply simp_all[2]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
apply (metis Abs_fresh_iff(3) atom_eq_iff finite_set fresh_Cons fresh_Nil fresh_atom fresh_eqvt_at fresh_finite_atom_set fresh_set lt.fsupp)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
apply (metis Abs_fresh_iff(3) atom_eq_iff finite_set fresh_Cons fresh_Nil fresh_atom fresh_eqvt_at fresh_finite_atom_set fresh_set lt.fsupp)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
--"-"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
apply (simp add: Abs1_eq(3))
2933
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
   144
apply (erule Abs_lst1_fcb2)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
   145
apply (simp_all add: Abs_fresh_iff fresh_Nil fresh_star_def eqvt_at_def)[4]
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
--"-"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
apply (rename_tac k' M N m' n')
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
apply (subgoal_tac "atom k \<sharp> CPS_sumC M \<and> atom k' \<sharp> CPS_sumC M \<and> atom k \<sharp> CPS_sumC N \<and> atom k' \<sharp> CPS_sumC N \<and>
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
                    atom m \<sharp> CPS_sumC N \<and> atom m' \<sharp> CPS_sumC N")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
prefer 2
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
apply (intro conjI)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
apply (erule fresh_eqvt_at, simp add: finite_supp, assumption)+
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
apply clarify
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
apply (case_tac "k = k'", case_tac [!] "m' = k",case_tac [!]"m = k'",case_tac[!] "m = m'")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
apply (simp_all add: Abs1_eq_iff lt.fresh flip_def[symmetric] fresh_at_base flip_fresh_fresh permute_eq_iff)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
by (metis flip_at_base_simps(3) flip_at_simps(2) flip_commute permute_flip_at)+
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   158
termination (eqvt) by (relation "measure size") (simp_all)
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
lemmas [simp] = fresh_Pair_elim CPS.simps(2,3)[simplified fresh_Pair_elim]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
lemma [simp]: "supp (M*) = supp M"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
  by (induct rule: CPS.induct, simp_all add: lt.supp supp_at_base fresh_at_base fresh_def supp_Pair)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
     (simp_all only: atom_eq_iff[symmetric], blast+)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
lemma [simp]: "x \<sharp> M* = x \<sharp> M"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
  unfolding fresh_def by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
nominal_primrec
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
  convert:: "lt => lt" ("_+" [250] 250)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
where
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
  "(Var x)+ = Var x"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
| "(Abs x M)+ = Abs x (M*)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
| "(M $ N)+ = M $ N"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
  unfolding convert_graph_def eqvt_def
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
  apply (rule, perm_simp, rule, rule)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
  apply (erule lt.exhaust)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
  apply (simp_all)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
  apply blast
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   180
  apply (simp add: Abs1_eq_iff CPS.eqvt)
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
  by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   183
termination (eqvt)
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   184
  by (relation "measure size") (simp_all)
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
lemma convert_supp[simp]:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
  shows "supp (M+) = supp M"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
  by (induct M rule: lt.induct, simp_all add: lt.supp)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
lemma convert_fresh[simp]:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
  shows "x \<sharp> (M+) = x \<sharp> M"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   192
  unfolding fresh_def by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
lemma [simp]:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  shows "isValue (p \<bullet> (M::lt)) = isValue M"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
  by (nominal_induct M rule: lt.strong_induct) auto
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   197
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
lemma [eqvt]:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   199
  shows "p \<bullet> isValue M = isValue (p \<bullet> M)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
  by (induct M rule: lt.induct) (perm_simp, rule refl)+
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   201
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
nominal_primrec
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
  Kapply :: "lt \<Rightarrow> lt \<Rightarrow> lt"       (infixl ";" 100)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
where
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   205
  "Kapply (Abs x M) K = K $ (Abs x M)+"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   206
| "Kapply (Var x) K = K $ Var x"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
| "isValue M \<Longrightarrow> isValue N \<Longrightarrow> Kapply (M $ N) K = M+ $ N+ $ K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   208
| "isValue M \<Longrightarrow> \<not>isValue N \<Longrightarrow> atom n \<sharp> M \<Longrightarrow> atom n \<sharp> K \<Longrightarrow>
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   209
    Kapply (M $ N) K = N; (Abs n (M+ $ Var n $ K))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   210
| "\<not>isValue M \<Longrightarrow> atom m \<sharp> N \<Longrightarrow> atom m \<sharp> K \<Longrightarrow> atom n \<sharp> m \<Longrightarrow> atom n \<sharp> K \<Longrightarrow>
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   211
    Kapply (M $ N) K = M; (Abs m (N*  $ (Abs n (Var m $ Var n $ K))))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   212
  unfolding Kapply_graph_def eqvt_def
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
  apply (rule, perm_simp, rule, rule)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   214
apply (simp_all)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
apply (case_tac x)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   216
apply (rule_tac y="a" in lt.exhaust)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   217
apply (auto)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   218
apply (case_tac "isValue lt1")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   219
apply (case_tac "isValue lt2")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   220
apply (auto)[1]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   221
apply (rule_tac x="(lt1, ba)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   222
apply (simp add: fresh_Pair_elim fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   223
apply (rule_tac x="(lt2, ba)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   224
apply (rule_tac x="(a, ba)" and ?'a="name" in obtain_fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   225
apply (simp add: fresh_Pair_elim fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   226
apply (auto simp add: Abs1_eq_iff eqvts)[1]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   227
apply (rename_tac M N u K)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   228
apply (subgoal_tac "Abs n (M+ $ n~ $ K) =  Abs u (M+ $ u~ $ K)")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   229
apply (simp only:)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   230
apply (auto simp add: Abs1_eq_iff flip_def[symmetric] lt.fresh fresh_at_base flip_fresh_fresh[symmetric])[1]
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   231
apply (subgoal_tac "Abs m (Na* $ Abs n (m~ $ n~ $ Ka)) = Abs ma (Na* $ Abs na (ma~ $ na~ $ Ka))")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   232
apply (simp only:)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   233
apply (simp add: Abs1_eq_iff flip_def[symmetric] lt.fresh fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   234
apply (subgoal_tac "Ka = (n \<leftrightarrow> na) \<bullet> Ka")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   235
apply (subgoal_tac "Ka = (m \<leftrightarrow> ma) \<bullet> Ka")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   236
apply (subgoal_tac "Ka = (n \<leftrightarrow> (m \<leftrightarrow> ma) \<bullet> na) \<bullet> Ka")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   237
apply (case_tac "m = ma")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   238
apply simp_all
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   239
apply rule
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   240
apply (auto simp add: flip_fresh_fresh[symmetric])
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   241
apply (metis flip_at_base_simps(3) flip_fresh_fresh permute_flip_at)+
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   242
done
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   243
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   244
termination (eqvt)
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   245
  by (relation "measure (\<lambda>(t, _). size t)") (simp_all)
2861
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   246
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   247
section{* lemma related to Kapply *}
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   248
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   249
lemma [simp]: "isValue V \<Longrightarrow> V; K = K $ (V+)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   250
  by (nominal_induct V rule: lt.strong_induct) auto
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   251
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   252
section{* lemma related to CPS conversion *}
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   253
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   254
lemma value_CPS:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   255
  assumes "isValue V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   256
  and "atom a \<sharp> V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   257
  shows "V* = Abs a (a~ $ V+)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   258
  using assms
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   259
proof (nominal_induct V avoiding: a rule: lt.strong_induct, simp_all add: lt.fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   260
  fix name :: name and lt aa
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   261
  assume a: "atom name \<sharp> aa" "\<And>b. \<lbrakk>isValue lt; atom b \<sharp> lt\<rbrakk> \<Longrightarrow> lt* = Abs b (b~ $ lt+)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   262
    "atom aa \<sharp> lt \<or> aa = name"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   263
  obtain ab :: name where b: "atom ab \<sharp> (name, lt, a)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   264
  show "Abs name lt* = Abs aa (aa~ $ Abs name (lt*))" using a b
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   265
    by (simp add: Abs1_eq_iff fresh_at_base lt.fresh)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   266
qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   267
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   268
section{* first lemma CPS subst *}
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   269
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   270
lemma CPS_subst_fv:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   271
  assumes *:"isValue V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   272
  shows "((M[V/x])* = (M*)[V+/x])"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   273
using * proof (nominal_induct M avoiding: V x rule: lt.strong_induct)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   274
  case (Var name)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   275
  assume *: "isValue V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   276
  obtain a :: name where a: "atom a \<sharp> (x, name, V)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   277
  show "((name~)[V/x])* = (name~)*[V+/x]" using a
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   278
    by (simp add: fresh_at_base * value_CPS)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   279
next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   280
  case (Abs name lt V x)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   281
  assume *: "atom name \<sharp> V" "atom name \<sharp> x" "\<And>b ba. isValue b \<Longrightarrow> (lt[b/ba])* = lt*[b+/ba]"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   282
    "isValue V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   283
  obtain a :: name where a: "atom a \<sharp> (name, lt, lt[V/x], x, V)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   284
  show "(Abs name lt[V/x])* = Abs name lt*[V+/x]" using * a
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   285
    by (simp add: fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   286
next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   287
  case (App lt1 lt2 V x)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   288
  assume *: "\<And>b ba. isValue b \<Longrightarrow> (lt1[b/ba])* = lt1*[b+/ba]" "\<And>b ba. isValue b \<Longrightarrow> (lt2[b/ba])* = lt2*[b+/ba]"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   289
    "isValue V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   290
  obtain a :: name where a: "atom a \<sharp> (lt1[V/x], lt1, lt2[V/x], lt2, V, x)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   291
  obtain b :: name where b: "atom b \<sharp> (lt2[V/x], lt2, a, V, x)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   292
  obtain c :: name where c: "atom c \<sharp> (a, b, V, x)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   293
  show "((lt1 $ lt2)[V/x])* = (lt1 $ lt2)*[V+/x]" using * a b c
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   294
    by (simp add: fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   295
qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   296
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   297
lemma [simp]: "isValue V \<Longrightarrow> isValue (V+)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   298
  by (nominal_induct V rule: lt.strong_induct, auto)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   299
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   300
lemma CPS_eval_Kapply:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   301
  assumes a: "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   302
  shows "(M* $ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* (M ; K)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   303
  using a
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   304
proof (nominal_induct M avoiding: K rule: lt.strong_induct, simp_all)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   305
  case (Var name K)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   306
  assume *: "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   307
  obtain a :: name where a: "atom a \<sharp> (name, K)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   308
  show "(name~)* $ K \<longrightarrow>\<^isub>\<beta>\<^sup>* K $ name~" using * a
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   309
    by simp (rule evbeta', simp_all add: fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   310
next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   311
  fix name :: name and lt K
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   312
  assume *: "atom name \<sharp> K" "\<And>b. isValue b \<Longrightarrow> lt* $ b \<longrightarrow>\<^isub>\<beta>\<^sup>* lt ; b" "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   313
  obtain a :: name where a: "atom a \<sharp> (name, K, lt)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   314
  then have b: "atom name \<sharp> a" using fresh_PairD(1) fresh_at_base atom_eq_iff by metis
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   315
  show "Abs name lt* $ K \<longrightarrow>\<^isub>\<beta>\<^sup>* K $ Abs name (lt*)" using * a b
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   316
    by simp (rule evbeta', simp_all)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   317
next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   318
  fix lt1 lt2 K
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   319
  assume *: "\<And>b. isValue b \<Longrightarrow>  lt1* $ b \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 ; b" "\<And>b. isValue b \<Longrightarrow>  lt2* $ b \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2 ; b" "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   320
  obtain a :: name where a: "atom a \<sharp> (lt1, lt2, K)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   321
  obtain b :: name where b: "atom b \<sharp> (lt1, lt2, K, a)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   322
  obtain c :: name where c: "atom c \<sharp> (lt1, lt2, K, a, b)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   323
  have d: "atom a \<sharp> lt1" "atom a \<sharp> lt2" "atom a \<sharp> K" "atom b \<sharp> lt1" "atom b \<sharp> lt2" "atom b \<sharp> K" "atom b \<sharp> a"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   324
    "atom c \<sharp> lt1" "atom c \<sharp> lt2" "atom c \<sharp> K" "atom c \<sharp> a" "atom c \<sharp> b" using fresh_Pair a b c by simp_all
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   325
  have "(lt1 $ lt2)* $ K \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1* $ Abs b (lt2* $ Abs c (b~ $ c~ $ K))" using * d
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   326
    by (simp add: fresh_at_base) (rule evbeta', simp_all add: fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   327
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 $ lt2 ; K" proof (cases "isValue lt1")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   328
    assume e: "isValue lt1"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   329
    have "lt1* $ Abs b (lt2* $ Abs c (b~ $ c~ $ K)) \<longrightarrow>\<^isub>\<beta>\<^sup>* Abs b (lt2* $ Abs c (b~ $ c~ $ K)) $ lt1+"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   330
      using * d e by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   331
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2* $ Abs c (lt1+ $ c~ $ K)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   332
      by (rule evbeta', simp_all add: * d e, metis d(12) fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   333
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 $ lt2 ; K" proof (cases "isValue lt2")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   334
      assume f: "isValue lt2"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   335
      have "lt2* $ Abs c (lt1+ $ c~ $ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* Abs c (lt1+ $ c~ $ K) $ lt2+" using * d e f by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   336
      also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1+ $ lt2+ $ K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   337
        by (rule evbeta', simp_all add: d e f)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   338
      finally show ?thesis using * d e f by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   339
    next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   340
      assume f: "\<not> isValue lt2"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   341
      have "lt2* $ Abs c (lt1+ $ c~ $ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2 ; Abs c (lt1+ $ c~ $ K)" using * d e f by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   342
      also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2 ; Abs a (lt1+ $ a~ $ K)" using Kapply.simps(4) d e evs1 f by metis
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   343
      finally show ?thesis using * d e f by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   344
    qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   345
    finally show ?thesis .
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   346
  qed (metis Kapply.simps(5) isValue.simps(2) * d)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   347
  finally show "(lt1 $ lt2)* $ K \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 $ lt2 ; K" .
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   348
qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   349
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   350
lemma Kapply_eval:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   351
  assumes a: "M \<longrightarrow>\<^isub>\<beta> N" "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   352
  shows "(M; K) \<longrightarrow>\<^isub>\<beta>\<^sup>*  (N; K)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   353
  using assms
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   354
proof (induct arbitrary: K rule: eval.induct)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   355
  case (evbeta x V M)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   356
  fix K
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   357
  assume a: "isValue K" "isValue V" "atom x \<sharp> V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   358
  have "Abs x (M*) $ V+ $ K \<longrightarrow>\<^isub>\<beta>\<^sup>* ((M*)[V+/x] $ K)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   359
    by (rule evs2,rule ev2,rule Lt.evbeta) (simp_all add: fresh_def a[simplified fresh_def] evs1)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   360
  also have "... = ((M[V/x])* $ K)" by (simp add: CPS_subst_fv a)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   361
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (M[V/x] ; K)" by (rule CPS_eval_Kapply, simp_all add: a)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   362
  finally show "(Abs x M $ V ; K) \<longrightarrow>\<^isub>\<beta>\<^sup>*  (M[V/x] ; K)" using a by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   363
next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   364
  case (ev1 V M N)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   365
  fix V M N K
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   366
  assume a: "isValue V" "M \<longrightarrow>\<^isub>\<beta> N" "\<And>K. isValue K \<Longrightarrow> M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* N ; K" "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   367
  obtain a :: name where b: "atom a \<sharp> (V, K, M, N)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   368
  show "V $ M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* V $ N ; K" proof (cases "isValue N")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   369
    assume "\<not> isValue N"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   370
    then show "V $ M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* V $ N ; K" using a b by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   371
  next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   372
    assume n: "isValue N"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   373
    have c: "M; Abs a (V+ $ a~ $ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* Abs a (V+ $ a~ $ K) $ N+" using a b by (simp add: n)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   374
    also have d: "... \<longrightarrow>\<^isub>\<beta>\<^sup>* V+ $ N+ $ K" by (rule evbeta') (simp_all add: a b n)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   375
    finally show "V $ M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* V $ N ; K" using a b by (simp add: n)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   376
  qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   377
next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   378
  case (ev2 M M' N)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   379
  assume *: "M \<longrightarrow>\<^isub>\<beta> M'" "\<And>K. isValue K \<Longrightarrow>  M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* M' ; K" "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   380
  obtain a :: name where a: "atom a \<sharp> (K, M, N, M')" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   381
  obtain b :: name where b: "atom b \<sharp> (a, K, M, N, M', N+)" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   382
  have d: "atom a \<sharp> K" "atom a \<sharp> M" "atom a \<sharp> N" "atom a \<sharp> M'" "atom b \<sharp> a" "atom b \<sharp> K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   383
    "atom b \<sharp> M" "atom b \<sharp> N" "atom b \<sharp> M'" using a b fresh_Pair by simp_all
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   384
  have "M $ N ; K  \<longrightarrow>\<^isub>\<beta>\<^sup>* M' ; Abs a (N* $ Abs b (a~ $ b~ $ K))" using * d by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   385
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* M' $ N ; K" proof (cases "isValue M'")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   386
    assume "\<not> isValue M'"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   387
    then show ?thesis using * d by (simp_all add: evs1)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   388
  next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   389
    assume e: "isValue M'"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   390
    then have "M' ; Abs a (N* $ Abs b (a~ $ b~ $ K)) = Abs a (N* $ Abs b (a~ $ b~ $ K)) $ M'+" by simp
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   391
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (N* $ Abs b (a~ $ b~ $ K))[M'+/a]"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   392
      by (rule evbeta') (simp_all add: fresh_at_base e d)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   393
    also have "... = N* $ Abs b (M'+ $ b~ $ K)" using * d by (simp add: fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   394
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* M' $ N ; K" proof (cases "isValue N")
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   395
      assume f: "isValue N"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   396
      have "N* $ Abs b (M'+ $ b~ $ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* Abs b (M'+ $ b~ $ K) $ N+"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   397
        by (rule eval_trans, rule CPS_eval_Kapply) (simp_all add: d e f * evs1)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   398
      also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* M' $ N ; K" by (rule evbeta') (simp_all add: d e f evs1)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   399
      finally show ?thesis .
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   400
    next
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   401
      assume "\<not> isValue N"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   402
      then show ?thesis using d e
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   403
        by (metis CPS_eval_Kapply Kapply.simps(4) isValue.simps(2))
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   404
    qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   405
    finally show ?thesis .
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   406
  qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   407
  finally show ?case .
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   408
qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   409
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   410
lemma Kapply_eval_rtrancl:
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   411
  assumes H: "M \<longrightarrow>\<^isub>\<beta>\<^sup>*  N" and "isValue K"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   412
  shows "(M;K) \<longrightarrow>\<^isub>\<beta>\<^sup>* (N;K)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   413
  using H
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   414
  by (induct) (metis Kapply_eval assms(2) eval_trans evs1)+
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   415
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   416
lemma
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   417
  assumes "isValue V" "M \<longrightarrow>\<^isub>\<beta>\<^sup>* V"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   418
  shows "M* $ (Abs x (x~)) \<longrightarrow>\<^isub>\<beta>\<^sup>* V+"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   419
proof-
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   420
  obtain y::name where *: "atom y \<sharp> V" using obtain_fresh by blast
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   421
  have e: "Abs x (x~) = Abs y (y~)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   422
    by (simp add: Abs1_eq_iff lt.fresh fresh_at_base)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   423
  have "M* $ Abs x (x~) \<longrightarrow>\<^isub>\<beta>\<^sup>* M ; Abs x (x~)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   424
    by(rule CPS_eval_Kapply,simp_all add: assms)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   425
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (V ; Abs x (x~))" by (rule Kapply_eval_rtrancl, simp_all add: assms)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   426
  also have "... = V ; Abs y (y~)" using e by (simp only:)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   427
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (V+)" by (simp add: assms, rule evbeta') (simp_all add: assms *)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   428
  finally show "M* $ (Abs x (x~)) \<longrightarrow>\<^isub>\<beta>\<^sup>* (V+)" .
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   429
qed
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   430
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   431
end
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   432
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   433
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   434