Nominal/Ex/Let.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 18 Jan 2011 21:26:58 +0100
changeset 2676 028d5511c15f
parent 2670 3c493c951388
child 2720 3dd6d524dfdd
permissions -rw-r--r--
some tryes about substitution over type-schemes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
     1
theory Let
2454
9ffee4eb1ae1 renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 2449
diff changeset
     2
imports "../Nominal2" 
1600
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
nominal_datatype trm =
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
     8
  Var "name"
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
     9
| App "trm" "trm"
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    10
| Lam x::"name" t::"trm"  bind  x in t
2490
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    11
| Let as::"assn" t::"trm"   bind "bn as" in t
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    12
and assn =
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    13
  ANil
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    14
| ACons "name" "trm" "assn"
1600
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
binder
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  bn
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
where
2490
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    18
  "bn ANil = []"
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    19
| "bn (ACons x t as) = (atom x) # (bn as)"
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    20
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    21
thm trm_assn.fv_defs
2492
5ac9a74d22fd post-processed eq_iff and supp threormes according to the fv-supp equality
Christian Urban <urbanc@in.tum.de>
parents: 2490
diff changeset
    22
thm trm_assn.eq_iff 
2490
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    23
thm trm_assn.bn_defs
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    24
thm trm_assn.perm_simps
2492
5ac9a74d22fd post-processed eq_iff and supp threormes according to the fv-supp equality
Christian Urban <urbanc@in.tum.de>
parents: 2490
diff changeset
    25
thm trm_assn.induct
5ac9a74d22fd post-processed eq_iff and supp threormes according to the fv-supp equality
Christian Urban <urbanc@in.tum.de>
parents: 2490
diff changeset
    26
thm trm_assn.inducts
2490
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    27
thm trm_assn.distinct
320775fa47ca some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    28
thm trm_assn.supp
2493
2e174807c891 added postprocessed fresh-lemmas for constructors
Christian Urban <urbanc@in.tum.de>
parents: 2492
diff changeset
    29
thm trm_assn.fresh
2617
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    30
thm trm_assn.exhaust
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    31
thm trm_assn.strong_exhaust
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    32
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    33
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    34
lemma lets_bla:
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    35
  "x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Let (ACons x (Var y) ANil) (Var x)) \<noteq> (Let (ACons x (Var z) ANil) (Var x))"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    36
  by (simp add: trm_assn.eq_iff)
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    37
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    38
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    39
lemma lets_ok:
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    40
  "(Let (ACons x (Var y) ANil) (Var x)) = (Let (ACons y (Var y) ANil) (Var y))"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    41
  apply (simp add: trm_assn.eq_iff Abs_eq_iff )
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    42
  apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    43
  apply (simp_all add: alphas atom_eqvt supp_at_base fresh_star_def trm_assn.bn_defs trm_assn.supp)
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    44
  done
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    45
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    46
lemma lets_ok3:
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    47
  "x \<noteq> y \<Longrightarrow>
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    48
   (Let (ACons x (App (Var y) (Var x)) (ACons y (Var y) ANil)) (App (Var x) (Var y))) \<noteq>
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    49
   (Let (ACons y (App (Var x) (Var y)) (ACons x (Var x) ANil)) (App (Var x) (Var y)))"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    50
  apply (simp add: trm_assn.eq_iff)
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    51
  done
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    52
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    53
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    54
lemma lets_not_ok1:
1685
721d92623c9d Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1658
diff changeset
    55
  "x \<noteq> y \<Longrightarrow>
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    56
   (Let (ACons x (Var x) (ACons y (Var y) ANil)) (App (Var x) (Var y))) \<noteq>
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    57
   (Let (ACons y (Var x) (ACons x (Var y) ANil)) (App (Var x) (Var y)))"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    58
  apply (simp add: alphas trm_assn.eq_iff trm_assn.supp fresh_star_def atom_eqvt Abs_eq_iff trm_assn.bn_defs)
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    59
  done
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    60
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    61
lemma lets_nok:
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    62
  "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow>
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    63
   (Let (ACons x (App (Var z) (Var z)) (ACons y (Var z) ANil)) (App (Var x) (Var y))) \<noteq>
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    64
   (Let (ACons y (Var z) (ACons x (App (Var z) (Var z)) ANil)) (App (Var x) (Var y)))"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    65
  apply (simp add: alphas trm_assn.eq_iff fresh_star_def trm_assn.bn_defs Abs_eq_iff trm_assn.supp trm_assn.distinct)
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    66
  done
2670
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    67
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    68
lemma
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    69
  fixes a b c :: name
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    70
  assumes x: "a \<noteq> c" and y: "b \<noteq> c"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    71
  shows "\<exists>p.([atom a], Var c) \<approx>lst (op =) supp p ([atom b], Var c)"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    72
  apply (rule_tac x="(a \<leftrightarrow> b)" in exI)
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    73
  apply (simp add: alphas trm_assn.supp supp_at_base x y fresh_star_def atom_eqvt)
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    74
  by (metis Rep_name_inverse atom_name_def flip_fresh_fresh fresh_atom fresh_perm x y)
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    75
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    76
lemma
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    77
  assumes "atom a \<noteq> atom c \<and> atom b \<noteq> atom c"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    78
  shows "\<exists>p. ([atom a], Var c) \<approx>lst (op =) supp p ([atom b], Var c) \<and> supp p \<subseteq> (supp (Var c) \<inter> {atom a}) \<union> (supp (Var c) \<inter> {atom b})"
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    79
  apply (simp add: alphas trm_assn.supp supp_at_base assms trm_assn.eq_iff atom_eqvt fresh_star_def)
3c493c951388 alpha_abs_let_stronger is not true in the same form
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2630
diff changeset
    80
  oops
1602
a7e60da429e2 Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1600
diff changeset
    81
1600
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
end
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
e33e37fd4c7d More reorganization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85