Nominal/Ex/CPS/CPS1_Plotkin.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 07 Aug 2012 18:54:52 +0100
changeset 3197 25d11b449e92
parent 3192 14c7d7e29c44
child 3231 188826f1ccdb
permissions -rw-r--r--
definition of an auxiliary graph in nominal-primrec definitions
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
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
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
     7
  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
     8
where
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
     9
  "atom k \<sharp> x \<Longrightarrow> (x~)* = (Lam k ((k~) $$ (x~)))"
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    10
| "atom k \<sharp> (x, M) \<Longrightarrow> (Lam x M)* = Lam k (k~ $$ Lam x (M*))"
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
    11
| "atom k \<sharp> (M, N) \<Longrightarrow> atom m \<sharp> (N, k) \<Longrightarrow> atom n \<sharp> (k, m) \<Longrightarrow>
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    12
    (M $$ N)* = Lam k (M* $$ Lam m (N* $$ Lam n (m~ $$ n~ $$ k~)))"
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
    13
unfolding eqvt_def CPS_graph_aux_def
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
    14
apply (simp)
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
    15
using [[simproc del: alpha_lst]]
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
    16
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
    17
apply (rule_tac y="x" in lt.exhaust)
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    18
apply (auto)[3]
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
    19
apply (rule_tac x="name" and ?'a="name" in obtain_fresh)
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
    20
using [[simproc del: alpha_lst]]
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
    21
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
    22
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
    23
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
    24
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
    25
apply (simp add: fresh_Pair_elim fresh_at_base)
3089
9bcf02a6eea9 Reorder constructors to match Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2998
diff changeset
    26
apply (rule_tac x="(name, lt)" and ?'a="name" in obtain_fresh)
9bcf02a6eea9 Reorder constructors to match Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2998
diff changeset
    27
apply (simp add: fresh_Pair_elim fresh_at_base)[2]
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
    28
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
    29
--"-"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    30
apply(rule_tac s="[[atom ka]]lst. ka~ $$ Lam x (CPS_sumC M)" in trans)
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
    31
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
    32
apply simp
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    33
thm Abs1_eq_iff
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    34
apply(subst Abs1_eq_iff)
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    35
apply(rule disjI2)
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    36
apply(rule conjI)
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    37
apply(simp)
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    38
apply(rule conjI)
3089
9bcf02a6eea9 Reorder constructors to match Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2998
diff changeset
    39
apply (simp only: lt.perm_simps(1) lt.perm_simps(2) flip_def[symmetric] lt.eq_iff)
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
    40
apply (subst  flip_at_base_simps(2))
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
    41
apply(simp)
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
    42
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
    43
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
    44
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
    45
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
    46
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
    47
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
    48
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
    49
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
    50
--"-"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
apply (simp add: Abs1_eq(3))
2933
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    52
apply (erule Abs_lst1_fcb2)
3be019a86117 Change CPS1 to FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2895
diff changeset
    53
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
    54
--"-"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
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
    56
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
    57
                    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
    58
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
    59
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
    60
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
    61
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
    62
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
    63
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
    64
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
    65
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
    66
termination (eqvt) by lexicographic_order
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
    67
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
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
    69
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
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
    71
  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
    72
     (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
    73
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
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
    75
  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
    76
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
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
    78
  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
    79
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
    80
  "(Var x)+ = Var x"
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
    81
| "(Lam x M)+ = Lam x (M*)"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    82
| "(M $$ N)+ = M $$ N"
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
    83
  unfolding convert_graph_aux_def eqvt_def
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
    84
  apply (simp)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
    85
  apply(rule TrueI)
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
    86
  apply (erule lt.exhaust)
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
    87
  using [[simproc del: alpha_lst]]
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
    88
  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
    89
  apply blast
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
    90
  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
    91
  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
    92
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
    93
termination (eqvt)
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
    94
  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
    95
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
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
    97
  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
    98
  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
    99
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
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
   101
  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
   102
  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
   103
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   105
  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
   106
  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
   107
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   109
  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
   110
where
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   111
  "Kapply (Lam x M) K = K $$ (Lam x M)+"
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   112
| "Kapply (Var x) K = K $$ Var x"
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   113
| "isValue M \<Longrightarrow> isValue N \<Longrightarrow> Kapply (M $$ N) K = M+ $$ N+ $$ K"
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
   114
| "isValue M \<Longrightarrow> \<not>isValue N \<Longrightarrow> atom n \<sharp> M \<Longrightarrow> atom n \<sharp> K \<Longrightarrow>
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   115
    Kapply (M $$ N) K = N; (Lam n (M+ $$ Var n $$ K))"
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
   116
| "\<not>isValue M \<Longrightarrow> atom m \<sharp> N \<Longrightarrow> atom m \<sharp> K \<Longrightarrow> atom n \<sharp> m \<Longrightarrow> atom n \<sharp> K \<Longrightarrow>
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   117
    Kapply (M $$ N) K = M; (Lam m (N*  $$ (Lam n (Var m $$ Var n $$ K))))"
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
   118
  unfolding Kapply_graph_aux_def eqvt_def
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3192
diff changeset
   119
  apply (simp)
3192
14c7d7e29c44 added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 3191
diff changeset
   120
using [[simproc del: alpha_lst]]
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
   121
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
   122
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
   123
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
   124
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
   125
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
   126
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
   127
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
   128
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
   129
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
   130
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
   131
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
   132
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
   133
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
   134
apply (rename_tac M N u K)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   135
apply (subgoal_tac "Lam n (M+ $$ n~ $$ K) =  Lam u (M+ $$ u~ $$ K)")
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
   136
apply (simp only:)
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3187
diff changeset
   137
apply (auto simp add: Abs1_eq_iff flip_fresh_fresh fresh_at_base)[1]
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   138
apply (subgoal_tac "Lam m (Na* $$ Lam n (m~ $$ n~ $$ Ka)) = Lam ma (Na* $$ Lam na (ma~ $$ na~ $$ Ka))")
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
   139
apply (simp only:)
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3187
diff changeset
   140
apply (simp add: Abs1_eq_iff flip_fresh_fresh fresh_at_base)
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
   141
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
   142
apply simp_all
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
   143
apply (case_tac "m = na")
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
   144
apply(simp_all add: flip_fresh_fresh)
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
   145
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
   146
2984
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   147
termination (eqvt)
1b39ba5db2c1 update to 'termination (eqvt)'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2933
diff changeset
   148
  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
   149
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   151
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   152
lemma [simp]: "isValue V \<Longrightarrow> V; K = K $$ (V+)"
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
   153
  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
   154
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   156
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   158
  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
   159
  and "atom a \<sharp> V"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   160
  shows "V* = Lam a (a~ $$ V+)"
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
   161
  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
   162
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
   163
  fix name :: name and lt aa
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   164
  assume a: "atom name \<sharp> aa" "\<And>b. \<lbrakk>isValue lt; atom b \<sharp> lt\<rbrakk> \<Longrightarrow> lt* = Lam b (b~ $$ lt+)"
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
   165
    "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
   166
  obtain ab :: name where b: "atom ab \<sharp> (name, lt, a)" using obtain_fresh by blast
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   167
  show "Lam name lt* = Lam aa (aa~ $$ Lam name (lt*))" using a b
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
   168
    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
   169
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
   170
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   172
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   174
  assumes *:"isValue V"
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   175
  shows "((M[x ::= V])* = (M*)[x ::= V+])"
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
   176
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
   177
  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
   178
  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
   179
  obtain a :: name where a: "atom a \<sharp> (x, name, V)" using obtain_fresh by blast
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   180
  show "((name~)[x ::= V])* = (name~)*[x ::= V+]" using a
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 (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
   182
next
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   183
  case (Lam name lt V x)
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   184
  assume *: "atom name \<sharp> V" "atom name \<sharp> x" "\<And>b ba. isValue b \<Longrightarrow> (lt[ba ::= b])* = lt*[ba ::= b+]"
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
    "isValue V"
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   186
  obtain a :: name where a: "atom a \<sharp> (name, lt, lt[x ::= V], x, V)" using obtain_fresh by blast
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   187
  show "(Lam name lt[x ::= V])* = Lam name lt*[x ::= V+]" using * a
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
   188
    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
   189
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
   190
  case (App lt1 lt2 V x)
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   191
  assume *: "\<And>b ba. isValue b \<Longrightarrow> (lt1[ba ::= b])* = lt1*[ba ::= b+]" "\<And>b ba. isValue b \<Longrightarrow> (lt2[ba ::= b])* = lt2*[ba ::= b+]"
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
   192
    "isValue V"
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   193
  obtain a :: name where a: "atom a \<sharp> (lt1[x ::= V], lt1, lt2[x ::= V], lt2, V, x)" using obtain_fresh by blast
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   194
  obtain b :: name where b: "atom b \<sharp> (lt2[x ::= V], lt2, a, V, x)" using obtain_fresh by blast
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
   195
  obtain c :: name where c: "atom c \<sharp> (a, b, V, x)" using obtain_fresh by blast
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   196
  show "((lt1 $$ lt2)[x ::= V])* = (lt1 $$ lt2)*[x ::= V+]" using * a b c
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
   197
    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
   198
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
   199
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   201
  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
   202
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   204
  assumes a: "isValue K"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   205
  shows "(M* $$ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* (M ; K)"
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
   206
  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
   207
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
   208
  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
   209
  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
   210
  obtain a :: name where a: "atom a \<sharp> (name, K)" using obtain_fresh by blast
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   211
  show "(name~)* $$ K \<longrightarrow>\<^isub>\<beta>\<^sup>* K $$ name~" using * a
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
   212
    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
   213
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
   214
  fix name :: name and lt K
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   215
  assume *: "atom name \<sharp> K" "\<And>b. isValue b \<Longrightarrow> lt* $$ b \<longrightarrow>\<^isub>\<beta>\<^sup>* lt ; b" "isValue K"
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
   216
  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
   217
  then have b: "atom name \<sharp> a" using fresh_PairD(1) fresh_at_base atom_eq_iff by metis
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   218
  show "Lam name lt* $$ K \<longrightarrow>\<^isub>\<beta>\<^sup>* K $$ Lam name (lt*)" using * a b
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
   219
    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
   220
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
   221
  fix lt1 lt2 K
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   222
  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"
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
   223
  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
   224
  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
   225
  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
   226
  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
   227
    "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
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   228
  have "(lt1 $$ lt2)* $$ K \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1* $$ Lam b (lt2* $$ Lam c (b~ $$ c~ $$ K))" using * d
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
   229
    by (simp add: fresh_at_base) (rule evbeta', simp_all add: fresh_at_base)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   230
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 $$ lt2 ; K" proof (cases "isValue lt1")
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
   231
    assume e: "isValue lt1"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   232
    have "lt1* $$ Lam b (lt2* $$ Lam c (b~ $$ c~ $$ K)) \<longrightarrow>\<^isub>\<beta>\<^sup>* Lam b (lt2* $$ Lam c (b~ $$ c~ $$ K)) $$ lt1+"
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
   233
      using * d e by simp
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   234
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2* $$ Lam c (lt1+ $$ c~ $$ K)"
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 3089
diff changeset
   235
      by (rule evbeta')(simp_all add: * d e)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   236
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 $$ lt2 ; K" proof (cases "isValue lt2")
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
   237
      assume f: "isValue lt2"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   238
      have "lt2* $$ Lam c (lt1+ $$ c~ $$ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* Lam c (lt1+ $$ c~ $$ K) $$ lt2+" using * d e f by simp
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   239
      also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1+ $$ lt2+ $$ K"
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
   240
        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
   241
      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
   242
    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
   243
      assume f: "\<not> isValue lt2"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   244
      have "lt2* $$ Lam c (lt1+ $$ c~ $$ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2 ; Lam c (lt1+ $$ c~ $$ K)" using * d e f by simp
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   245
      also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* lt2 ; Lam a (lt1+ $$ a~ $$ K)" using Kapply.simps(4) d e evs1 f by metis
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
      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
   247
    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
   248
    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
   249
  qed (metis Kapply.simps(5) isValue.simps(2) * d)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   250
  finally show "(lt1 $$ lt2)* $$ K \<longrightarrow>\<^isub>\<beta>\<^sup>* lt1 $$ lt2 ; K" .
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
   251
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
   252
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   254
  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
   255
  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
   256
  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
   257
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
   258
  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
   259
  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
   260
  assume a: "isValue K" "isValue V" "atom x \<sharp> V"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   261
  have "Lam x (M*) $$ V+ $$ K \<longrightarrow>\<^isub>\<beta>\<^sup>* (((M*)[x ::= V+]) $$ K)"
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
   262
    by (rule evs2,rule ev2,rule Lt.evbeta) (simp_all add: fresh_def a[simplified fresh_def] evs1)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   263
  also have "... = ((M[x ::= V])* $$ K)" by (simp add: CPS_subst_fv a)
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   264
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* ((M[x ::= V]) ; K)" by (rule CPS_eval_Kapply, simp_all add: a)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   265
  finally show "(Lam x M $$ V ; K) \<longrightarrow>\<^isub>\<beta>\<^sup>*  ((M[x ::= V]) ; K)" using a by simp
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
   266
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
   267
  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
   268
  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
   269
  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
   270
  obtain a :: name where b: "atom a \<sharp> (V, K, M, N)" using obtain_fresh by blast
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   271
  show "V $$ M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* V $$ N ; K" proof (cases "isValue N")
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
   272
    assume "\<not> isValue N"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   273
    then show "V $$ M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* V $$ N ; K" using a b by simp
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
   274
  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
   275
    assume n: "isValue N"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   276
    have c: "M; Lam a (V+ $$ a~ $$ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* Lam a (V+ $$ a~ $$ K) $$ N+" using a b by (simp add: n)
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   277
    also have d: "... \<longrightarrow>\<^isub>\<beta>\<^sup>* V+ $$ N+ $$ K" by (rule evbeta') (simp_all add: a b n)
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   278
    finally show "V $$ M ; K \<longrightarrow>\<^isub>\<beta>\<^sup>* V $$ N ; K" using a b by (simp add: n)
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
   279
  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
   280
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
   281
  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
   282
  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
   283
  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
   284
  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
   285
  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
   286
    "atom b \<sharp> M" "atom b \<sharp> N" "atom b \<sharp> M'" using a b fresh_Pair by simp_all
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   287
  have "M $$ N ; K  \<longrightarrow>\<^isub>\<beta>\<^sup>* M' ; Lam a (N* $$ Lam b (a~ $$ b~ $$ K))" using * d by simp
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   288
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* M' $$ N ; K" proof (cases "isValue M'")
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
   289
    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
   290
    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
   291
  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
   292
    assume e: "isValue M'"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   293
    then have "M' ; Lam a (N* $$ Lam b (a~ $$ b~ $$ K)) = Lam a (N* $$ Lam b (a~ $$ b~ $$ K)) $$ M'+" by simp
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   294
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (N* $$ Lam b (a~ $$ b~ $$ K))[a ::= M'+]"
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
   295
      by (rule evbeta') (simp_all add: fresh_at_base e d)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   296
    also have "... = N* $$ Lam b (M'+ $$ b~ $$ K)" using * d by (simp add: fresh_at_base)
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   297
    also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* M' $$ N ; K" proof (cases "isValue N")
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
   298
      assume f: "isValue N"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   299
      have "N* $$ Lam b (M'+ $$ b~ $$ K) \<longrightarrow>\<^isub>\<beta>\<^sup>* Lam b (M'+ $$ b~ $$ K) $$ N+"
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
   300
        by (rule eval_trans, rule CPS_eval_Kapply) (simp_all add: d e f * evs1)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   301
      also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* M' $$ N ; K" by (rule evbeta') (simp_all add: d e f evs1)
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
   302
      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
   303
    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
   304
      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
   305
      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
   306
        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
   307
    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
   308
    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
   309
  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
   310
  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
   311
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
   312
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   314
  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
   315
  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
   316
  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
   317
  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
   318
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   320
  assumes "isValue V" "M \<longrightarrow>\<^isub>\<beta>\<^sup>* V"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   321
  shows "M* $$ (Lam x (x~)) \<longrightarrow>\<^isub>\<beta>\<^sup>* V+"
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
   322
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
   323
  obtain y::name where *: "atom y \<sharp> V" using obtain_fresh by blast
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   324
  have e: "Lam x (x~) = Lam y (y~)"
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
   325
    by (simp add: Abs1_eq_iff lt.fresh fresh_at_base)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   326
  have "M* $$ Lam x (x~) \<longrightarrow>\<^isub>\<beta>\<^sup>* M ; Lam x (x~)"
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
   327
    by(rule CPS_eval_Kapply,simp_all add: assms)
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   328
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (V ; Lam x (x~))" by (rule Kapply_eval_rtrancl, simp_all add: assms)
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2984
diff changeset
   329
  also have "... = V ; Lam y (y~)" using e by (simp only:)
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
   330
  also have "... \<longrightarrow>\<^isub>\<beta>\<^sup>* (V+)" by (simp add: assms, rule evbeta') (simp_all add: assms *)
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
   331
  finally show "M* $$ (Lam x (x~)) \<longrightarrow>\<^isub>\<beta>\<^sup>* (V+)" .
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
   332
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
   333
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
   335
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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