Nominal/Ex/CPS/CPS3_DanvyFilinski_FCB2.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 07 Aug 2012 18:54:52 +0100
changeset 3197 25d11b449e92
parent 3187 b3d97424b130
child 3235 5ebd327ffb96
permissions -rw-r--r--
definition of an auxiliary graph in nominal-primrec definitions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2934
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
header {* CPS transformation of Danvy and Filinski *}
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 2998
diff changeset
     2
theory CPS3_DanvyFilinski_FCB2 
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 2998
diff changeset
     3
imports Lt 
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 2998
diff changeset
     4
begin
2934
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
nominal_primrec
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
  CPS1 :: "lt \<Rightarrow> (lt \<Rightarrow> lt) \<Rightarrow> lt" ("_*_"  [100,100] 100)
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
and
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
  CPS2 :: "lt \<Rightarrow> lt \<Rightarrow> lt" ("_^_" [100,100] 100)
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
where
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
  "eqvt k \<Longrightarrow> (x~)*k = k (x~)"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    12
| "eqvt k \<Longrightarrow> (M $$ N)*k = M*(%m. (N*(%n.((m $$ n) $$ (Lam c (k (c~)))))))"
2998
f0fab367453a Use same constructor names as Lambda, remove copies of FCB, remove [eqvt].
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2965
diff changeset
    13
| "eqvt k \<Longrightarrow> atom c \<sharp> (x, M) \<Longrightarrow> (Lam x M)*k = k (Lam x (Lam c (M^(c~))))"
2934
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
| "\<not>eqvt k \<Longrightarrow> (CPS1 t k) = t"
3187
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    15
| "(x~)^l = l $$ (x~)"
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    16
| "(M $$ N)^l = M*(%m. (N*(%n.((m $$ n) $$ l))))"
b3d97424b130 added finfun-type to Nominal
Christian Urban <urbanc@in.tum.de>
parents: 3186
diff changeset
    17
| "atom c \<sharp> (x, M) \<Longrightarrow> (Lam x M)^l = l $$ (Lam x (Lam c (M^(c~))))"
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3187
diff changeset
    18
  apply (simp add: eqvt_def CPS1_CPS2_graph_aux_def)
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 2998
diff changeset
    19
  apply (auto simp only:)
2934
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
  apply (case_tac x)
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
  apply (case_tac a)
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
  apply (case_tac "eqvt b")
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  apply (rule_tac y="aa" in lt.strong_exhaust)
3186
425b4c406d80 added CPS files to test (not all proofs have been completed)
Christian Urban <urbanc@in.tum.de>
parents: 2998
diff changeset
    24
  apply auto
2964
0d95e19e4f93 Remove copy of FCB and cleanup
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2934
diff changeset
    25
  oops
2934
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
end
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
78fc2bd14d02 Added non-working CPS3 using FCB2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30