Nominal/Ex/CPS/CPS2_DanvyNielsen.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Sun, 19 Jun 2011 13:10:15 +0900
changeset 2869 9c0df9901acf
parent 2864 bb647489f130
child 2895 65efa1c7563c
permissions -rw-r--r--
little on cps2
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 transformation of Danvy and Nielsen *}
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
theory DanvyNielsen
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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_datatype cpsctxt =
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
  Hole
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
| CFun cpsctxt 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
     9
| CArg lt cpsctxt
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
| CAbs x::name c::cpsctxt bind x in 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
    11
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
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
    13
  fill   :: "cpsctxt \<Rightarrow> lt \<Rightarrow> lt"         ("_<_>" [200, 200] 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
    14
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
    15
  fill_hole : "Hole<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
    16
| fill_fun  : "(CFun C N)<M> = (C<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
    17
| fill_arg  : "(CArg N C)<M> = N $ (C<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
    18
| fill_abs  : "atom x \<sharp> M \<Longrightarrow> (CAbs x C)<M> = Abs x (C<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
    19
  unfolding eqvt_def fill_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
    20
  apply perm_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
    21
  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
    22
  apply (rule_tac y="a" and c="b" in cpsctxt.strong_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
    23
  apply (auto simp add: fresh_star_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
    24
  apply (erule Abs_lst1_fcb)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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_all add: Abs_fresh_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
    26
  apply (erule fresh_eqvt_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
    27
  apply (simp add: finite_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
    28
  apply (simp add: fresh_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
    29
  apply (simp add: eqvt_at_def swap_fresh_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
    30
  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
    31
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
termination
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
  by (relation "measure (\<lambda>(x, _). size x)") (auto simp add: cpsctxt.size)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
lemma [eqvt]: "p \<bullet> fill c t = fill (p \<bullet> c) (p \<bullet> t)"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
  by (nominal_induct c avoiding: t rule: cpsctxt.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
    37
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
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
    39
  ccomp :: "cpsctxt => cpsctxt => cpsctxt"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
    41
  "ccomp Hole C  = 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
    42
| "atom x \<sharp> C' \<Longrightarrow> ccomp (CAbs x C) C' = CAbs x (ccomp C 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
    43
| "ccomp (CArg N C) C' = CArg N (ccomp C 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
    44
| "ccomp (CFun C N) C'  = CFun (ccomp C C') 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
    45
  unfolding eqvt_def ccomp_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
    46
  apply perm_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
    47
  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
    48
  apply (rule_tac y="a" and c="b" in cpsctxt.strong_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
    49
  apply (auto simp add: fresh_star_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
    50
  apply 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
    51
  apply (erule Abs_lst1_fcb)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
  apply (simp_all add: Abs_fresh_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
    53
  apply (erule fresh_eqvt_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
    54
  apply (simp add: finite_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
    55
  apply (simp add: fresh_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
    56
  apply (simp add: eqvt_at_def swap_fresh_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
    57
  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
    58
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
termination
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
  by (relation "measure (\<lambda>(x, _). size x)") (auto simp add: cpsctxt.size)
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
lemma [eqvt]: "p \<bullet> ccomp c c' = ccomp (p \<bullet> c) (p \<bullet> 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
    63
  by (nominal_induct c avoiding: c' rule: cpsctxt.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
    64
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
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
    66
    CPSv :: "lt => 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
    67
and CPS :: "lt => cpsctxt" 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
    68
  "CPSv (Var 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
    69
| "CPS (Var x) = CFun Hole (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
    70
| "atom b \<sharp> M \<Longrightarrow> CPSv (Abs y M) = Abs y (Abs b ((CPS M)<Var 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
    71
| "atom b \<sharp> M \<Longrightarrow> CPS (Abs y M) = CFun Hole (Abs y (Abs b ((CPS M)<Var 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
    72
| "CPSv (M $ N) = Abs 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
    73
| "isValue M \<Longrightarrow> isValue N \<Longrightarrow> CPS (M $ N) = CArg (CPSv M $ CPSv N) Hole"
2869
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    74
| "isValue M \<Longrightarrow> ~isValue N \<Longrightarrow> atom a \<sharp> M \<Longrightarrow> CPS (M $ 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
    75
     ccomp (CPS N) (CAbs a (CArg (CPSv M $ Var a) Hole))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
| "~isValue M \<Longrightarrow> isValue N \<Longrightarrow> atom a \<sharp> N \<Longrightarrow> CPS (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
    77
     ccomp (CPS M) (CAbs a (CArg (Var a $ (CPSv N)) Hole))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
| "~isValue M \<Longrightarrow> ~isValue N \<Longrightarrow> atom a \<sharp> (N, b) \<Longrightarrow> CPS (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
    79
     ccomp (CPS M) (CAbs a (ccomp (CPS N) (CAbs b (CArg (Var a $ Var b) Hole))))"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
  apply auto
2869
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    81
  defer
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    82
  apply (case_tac x)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    83
  apply (rule_tac y="a" in lt.exhaust)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    84
  apply auto
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    85
  apply (rule_tac x="lt" and ?'a="name" in obtain_fresh)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    86
  apply (simp add: Abs1_eq_iff)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    87
  apply blast+
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    88
  apply (rule_tac y="b" in lt.exhaust)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    89
  apply auto
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    90
  apply (rule_tac x="lt" and ?'a="name" in obtain_fresh)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    91
  apply (simp add: Abs1_eq_iff)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    92
  apply blast
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    93
  apply (rule_tac ?'a="name" in obtain_fresh)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    94
  apply (rule_tac x="(lt1, lt2, a)" and ?'a="name" in obtain_fresh)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    95
  apply (simp add: fresh_Pair_elim)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    96
  apply (case_tac "isValue lt1", case_tac [!] "isValue lt2")[1]
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    97
  apply (simp_all add: fresh_Pair)[4]
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    98
--""
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
    99
  apply (rule_tac x="(y, ya, b, ba, M, Ma)" and ?'a="name" in obtain_fresh)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   100
  apply (subgoal_tac "Abs b (Sum_Type.Projr (CPSv_CPS_sumC (Inr M))<(b~)>) = Abs a (Sum_Type.Projr (CPSv_CPS_sumC (Inr M))<(a~)>)")
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   101
  apply (subgoal_tac "Abs ba (Sum_Type.Projr (CPSv_CPS_sumC (Inr Ma))<(ba~)>) = Abs a (Sum_Type.Projr (CPSv_CPS_sumC (Inr Ma))<(a~)>)")
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   102
  apply (simp only:)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   103
  apply (erule Abs_lst1_fcb)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   104
  apply (simp add: Abs_fresh_iff)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   105
  apply (simp_all add: Abs_fresh_iff lt.fresh fresh_Pair_elim fresh_at_base swap_fresh_fresh)[2]
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   106
  (* need an invariant to get eqvt_at (Proj) *)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   107
  defer defer
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   108
  apply simp
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   109
  apply (simp_all add: Abs1_eq_iff fresh_Pair_elim fresh_at_base)[2]
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   110
  defer defer
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   111
  defer
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   112
  apply (simp add: Abs1_eq_iff fresh_at_base lt.fresh)
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   113
  apply (rule arg_cong) back
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   114
  defer
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   115
  apply (rule arg_cong) back
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   116
  defer
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   117
  apply (rule arg_cong) back
9c0df9901acf little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2864
diff changeset
   118
  defer
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
   119
  oops --"The goals seem reasonable"
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
5635a968fd3f Added the CPS translation experiments. CPS1 comes with all the 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
end