Nominal/Ex/NBE.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 06 Jul 2011 23:11:30 +0200
changeset 2956 7e1c309bf820
parent 2955 4049a2651dd9
child 2967 d7e8b9b78e28
permissions -rw-r--r--
more on NBE
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Lambda
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports 
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
  "../Nominal2"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
atom_decl name
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
nominal_datatype lam =
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  Var "name"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
| App "lam" "lam"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
| Lam x::"name" l::"lam"  binds x in l ("Lam [_]. _" [100, 100] 100)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
nominal_datatype sem =
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
    16
  L e::"env" x::"name" l::"lam" binds x "bn e" in l
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
| N "neu"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
and neu = 
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  V "name"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
| A "neu" "sem"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
and env =
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  ENil
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
| ECons "env" "name" "sem"
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    24
binder
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    25
  bn
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    26
where
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    27
  "bn ENil = []"
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
    28
| "bn (ECons env x v) = (atom x) # (bn env)" 
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
    29
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
2953
80f01215d1a6 Setup eqvt_at for first goal
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2952
diff changeset
    31
nominal_primrec
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  evals :: "env \<Rightarrow> lam \<Rightarrow> sem" and
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  evals_aux :: "sem \<Rightarrow> lam \<Rightarrow> env \<Rightarrow> sem"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
where
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  "evals ENil (Var x) = N (V x)"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
| "evals (ECons tail y v) (Var x) = (if x = y then v else evals tail (Var x))" 
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
| "evals env (Lam [x]. t) = L env x t"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
| "evals env (App t1 t2) = evals_aux (evals env t1) t2 env"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
| "evals_aux (L cenv x t) t2 env = evals (ECons cenv x (evals env t2)) t"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
| "evals_aux (N n) t2 env = N (A n (evals env t2))"
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    41
apply(subgoal_tac "\<And>p x y. evals_evals_aux_graph x y \<Longrightarrow> evals_evals_aux_graph (p \<bullet> x) (p \<bullet> y)")
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    42
apply(simp add: eqvt_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    43
apply(simp add: permute_fun_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    44
apply(rule allI)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    45
apply(rule ext)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    46
apply(rule ext)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    47
apply(rule iffI)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    48
apply(drule_tac x="p" in meta_spec)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    49
apply(drule_tac x="- p \<bullet> x" in meta_spec)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    50
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    51
apply(simp add: permute_bool_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    52
apply(simp add: permute_bool_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    53
apply(erule evals_evals_aux_graph.induct)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    54
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    55
apply(rule evals_evals_aux_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    56
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    57
apply(rule evals_evals_aux_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    58
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    59
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    60
apply(rule evals_evals_aux_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    61
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    62
apply(rule evals_evals_aux_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    63
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    64
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    65
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    66
apply(rule evals_evals_aux_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    67
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    68
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    69
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    70
apply(rule evals_evals_aux_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    71
apply(simp)
2953
80f01215d1a6 Setup eqvt_at for first goal
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2952
diff changeset
    72
apply (rule TrueI)
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
--"completeness"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
apply(case_tac x)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
apply(case_tac a)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
apply(case_tac aa rule: sem_neu_env.exhaust(3))
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
apply(case_tac b rule: lam.exhaust)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
apply(metis)+
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
apply(case_tac b rule: lam.exhaust)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
apply(metis)+
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
apply(case_tac b)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
apply(case_tac a rule: sem_neu_env.exhaust(1))
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
apply(metis)+
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
--"compatibility"
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
apply(all_trivials)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
apply(simp)
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    93
defer
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
apply(simp)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
apply(simp)
2953
80f01215d1a6 Setup eqvt_at for first goal
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2952
diff changeset
    96
apply (simp add: meta_eq_to_obj_eq[OF evals_def, symmetric, unfolded fun_eq_iff])
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
    97
apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons cenv x (evals enva t2a), t)")
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
    98
apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (enva, t2a)")
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
    99
apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons cenva xa (evals enva t2a), ta)")
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   100
apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons cenv x (evals enva t2a), t))")
2953
80f01215d1a6 Setup eqvt_at for first goal
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2952
diff changeset
   101
apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (enva, t2a))")
80f01215d1a6 Setup eqvt_at for first goal
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2952
diff changeset
   102
apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons cenva xa (evals enva t2a), ta))")
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   103
apply(erule conjE)+
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
defer
2953
80f01215d1a6 Setup eqvt_at for first goal
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2952
diff changeset
   105
apply (simp_all add: eqvt_at_def evals_def)[3]
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   106
apply(simp)
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
sorry
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
(* can probably not proved by a trivial size argument *)
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
termination sorry
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   112
lemma [eqvt]:
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   113
  shows "(p \<bullet> evals env t) = evals (p \<bullet> env) (p \<bullet> t)"
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   114
  and "(p \<bullet> evals_aux v t env) = evals_aux (p \<bullet> v) (p \<bullet> t) (p \<bullet> env)"
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   115
sorry
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   116
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   117
(* fixme: should be a provided lemma *)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   118
lemma fv_bn_finite:
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   119
  shows "finite (fv_bn env)"
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   120
apply(induct env rule: sem_neu_env.inducts(3))
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   121
apply(auto simp add: sem_neu_env.supp finite_supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   122
done
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   123
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   124
lemma test:
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   125
  fixes env::"env"
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   126
  shows "supp (evals env t) \<subseteq> (supp t - set (bn env)) \<union> (fv_bn env)"
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   127
  and "supp (evals_aux s t env) \<subseteq> (supp s) \<union> (fv_bn env) \<union> (supp (t) - set (bn env))"
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   128
apply(induct env t and s t env rule: evals_evals_aux.induct)
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   129
apply(simp add: sem_neu_env.supp lam.supp supp_Nil sem_neu_env.bn_defs)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   130
apply(simp add: sem_neu_env.supp lam.supp supp_Nil supp_Cons sem_neu_env.bn_defs)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   131
apply(rule conjI)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   132
apply(auto)[1]
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   133
apply(rule impI)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   134
apply(simp)
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   135
apply(simp add: supp_at_base)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   136
apply(blast)
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   137
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   138
apply(subst sem_neu_env.supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   139
apply(simp add: sem_neu_env.supp lam.supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   140
apply(auto)[1]
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   141
apply(simp add: lam.supp sem_neu_env.supp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   142
apply(blast)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   143
apply(simp add: sem_neu_env.supp sem_neu_env.bn_defs)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   144
prefer 2
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   145
apply(simp add: sem_neu_env.supp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   146
apply(rule conjI)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   147
apply(blast)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   148
apply(blast)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   149
apply(blast)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   150
done
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   151
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   152
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   153
nominal_primrec
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   154
  reify :: "sem \<Rightarrow> lam" and
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   155
  reifyn :: "neu \<Rightarrow> lam"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   156
where
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   157
  "atom x \<sharp> (env, y, t) \<Longrightarrow> reify (L env y t) = Lam [x]. (reify (evals (ECons env y (N (V x))) t))"
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   158
| "reify (N n) = reifyn n"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   159
| "reifyn (V x) = Var x"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   160
| "reifyn (A n d) = App (reifyn n) (reify d)"
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   161
apply(subgoal_tac "\<And>p x y. reify_reifyn_graph x y \<Longrightarrow> reify_reifyn_graph (p \<bullet> x) (p \<bullet> y)")
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   162
apply(simp add: eqvt_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   163
apply(simp add: permute_fun_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   164
apply(rule allI)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   165
apply(rule ext)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   166
apply(rule ext)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   167
apply(rule iffI)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   168
apply(drule_tac x="p" in meta_spec)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   169
apply(drule_tac x="- p \<bullet> x" in meta_spec)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   170
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   171
apply(simp add: permute_bool_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   172
apply(simp add: permute_bool_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   173
apply(erule reify_reifyn_graph.induct)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   174
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   175
apply(rule reify_reifyn_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   176
apply(rule_tac p="-p" in permute_boolE)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   177
apply(perm_simp add: permute_minus_cancel)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   178
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   179
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   180
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   181
apply(rule reify_reifyn_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   182
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   183
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   184
apply(rule reify_reifyn_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   185
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   186
apply(rule reify_reifyn_graph.intros)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   187
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   188
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   189
apply(rule TrueI)
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   190
--"completeness"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   191
apply(case_tac x)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   192
apply(simp)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   193
apply(case_tac a rule: sem_neu_env.exhaust(1))
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   194
apply(subgoal_tac "\<exists>x::name. atom x \<sharp> (env, name, lam)")
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   195
apply(metis)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   196
apply(rule obtain_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   197
apply(blast)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   198
apply(blast)
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   199
apply(case_tac b rule: sem_neu_env.exhaust(2))
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   200
apply(simp)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   201
apply(simp)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   202
apply(metis)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   203
--"compatibility"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   204
apply(all_trivials)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   205
defer
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   206
apply(simp)
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   207
apply(simp)
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   208
apply(simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   209
apply(erule conjE)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   210
apply (simp add: meta_eq_to_obj_eq[OF reify_def, symmetric, unfolded fun_eq_iff])
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   211
apply (subgoal_tac "eqvt_at (\<lambda>t. reify t) (evals (ECons env y (N (V x))) t)")
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   212
apply (subgoal_tac "eqvt_at (\<lambda>t. reify t) (evals (ECons enva ya (N (V xa))) ta)")
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   213
apply (thin_tac "eqvt_at reify_reifyn_sumC (Inl (evals (ECons env y (N (V x))) t))")
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   214
apply (thin_tac "eqvt_at reify_reifyn_sumC (Inl (evals (ECons enva ya (N (V xa))) ta))")
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   215
defer
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   216
apply (simp_all add: eqvt_at_def reify_def)[2]
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   217
apply(subgoal_tac "\<exists>c::name. atom c \<sharp> (x, xa, env, enva, y, ya, t, ta)")
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   218
prefer 2
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   219
apply(rule obtain_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   220
apply(blast)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   221
apply(erule exE)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   222
apply(rule trans)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   223
apply(rule sym)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   224
apply(rule_tac a="x" and b="c" in flip_fresh_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   225
apply(simp add: Abs_fresh_iff)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   226
apply(simp add: Abs_fresh_iff fresh_Pair)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   227
apply(auto)[1]
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   228
apply(rule fresh_eqvt_at)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   229
back
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   230
apply(assumption)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   231
apply(simp add: finite_supp)
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   232
apply(rule_tac S="supp (env, y, x, t)" in supports_fresh)
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   233
apply(simp add: supports_def fresh_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   234
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   235
apply(simp add: swap_fresh_fresh fresh_Pair)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   236
apply(simp add: finite_supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   237
apply(simp add: fresh_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   238
apply(simp add: eqvt_at_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   239
apply(simp add: eqvt_at_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   240
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   241
apply(simp add: flip_fresh_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   242
apply(rule sym)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   243
apply(rule trans)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   244
apply(rule sym)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   245
apply(rule_tac a="xa" and b="c" in flip_fresh_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   246
apply(simp add: Abs_fresh_iff)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   247
apply(simp add: Abs_fresh_iff fresh_Pair)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   248
apply(auto)[1]
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   249
apply(rule fresh_eqvt_at)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   250
back
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   251
apply(assumption)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   252
apply(simp add: finite_supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   253
apply(rule_tac S="supp (enva, ya, xa, ta)" in supports_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   254
apply(simp add: supports_def fresh_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   255
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   256
apply(simp add: swap_fresh_fresh fresh_Pair)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   257
apply(simp add: finite_supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   258
apply(simp add: fresh_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   259
apply(simp add: eqvt_at_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   260
apply(simp add: eqvt_at_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   261
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   262
apply(simp add: flip_fresh_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   263
apply(simp (no_asm) add: Abs1_eq_iff)
2956
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   264
thm at_set_avoiding3
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   265
using at_set_avoiding3
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   266
apply -
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   267
apply(drule_tac x="set (atom y # bn env)" in meta_spec)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   268
apply(drule_tac x="(env, enva)" in meta_spec)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   269
apply(drule_tac x="[atom y # bn env]lst. t" in meta_spec)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   270
apply(simp (no_asm_use) add: finite_supp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   271
apply(drule meta_mp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   272
apply(rule Abs_fresh_star)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   273
apply(auto)[1]
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   274
apply(erule exE)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   275
apply(erule conjE)+
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   276
apply(drule_tac q="(x \<leftrightarrow> c)" in eqvt_at_perm)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   277
apply(perm_simp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   278
apply(simp add: flip_fresh_fresh fresh_Pair)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   279
apply(drule_tac q="(xa \<leftrightarrow> c)" in eqvt_at_perm)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   280
apply(perm_simp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   281
apply(simp add: flip_fresh_fresh fresh_Pair)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   282
apply(drule sym)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   283
apply(rotate_tac 9)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   284
apply(drule trans)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   285
apply(rule sym)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   286
(* HERE *)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   287
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   288
apply(rule_tac p="p" in supp_perm_eq)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   289
apply(simp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   290
apply(perm_simp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   291
apply(simp add: Abs_eq_iff2 alphas)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   292
apply(clarify)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   293
apply(rule trans)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   294
apply(rule sym)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   295
apply(rule_tac p="pa" in perm_supp_eq)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   296
defer
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   297
apply(rule sym)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   298
apply(rule trans)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   299
apply(rule sym)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   300
apply(rule_tac p="p" in perm_supp_eq)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   301
defer
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   302
apply(simp add: atom_eqvt)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   303
apply(drule_tac q="(x \<leftrightarrow> c)" in eqvt_at_perm)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   304
apply(perm_simp)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   305
apply(simp add: flip_fresh_fresh fresh_Pair)
7e1c309bf820 more on NBE
Christian Urban <urbanc@in.tum.de>
parents: 2955
diff changeset
   306
2955
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   307
apply(rule sym)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   308
apply(erule_tac Abs_lst1_fcb2')
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   309
apply(rule fresh_eqvt_at)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   310
back
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   311
apply(drule_tac q="(c \<leftrightarrow> x)" in eqvt_at_perm)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   312
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   313
apply(simp add: flip_fresh_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   314
apply(simp add: finite_supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   315
apply(rule supports_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   316
apply(rule_tac S="supp (enva, ya, xa, ta)" in supports_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   317
apply(simp add: supports_def fresh_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   318
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   319
apply(simp add: swap_fresh_fresh fresh_Pair)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   320
apply(simp add: finite_supp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   321
apply(simp add: fresh_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   322
apply(simp add: eqvt_at_def)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   323
apply(simp add: eqvt_at_def[symmetric])
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   324
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   325
apply(rule fresh_eqvt_at)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   326
back
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   327
apply(drule_tac q="(c \<leftrightarrow> x)" in eqvt_at_perm)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   328
apply(perm_simp)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   329
apply(simp add: flip_fresh_fresh)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   330
apply(assumption)
4049a2651dd9 more on the NBE function
Christian Urban <urbanc@in.tum.de>
parents: 2954
diff changeset
   331
apply(simp add: finite_supp)
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   332
sorry
2952
e4c2854833ad attempt for NBE
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
2954
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   334
termination sorry
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   335
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   336
definition
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   337
  eval :: "lam \<Rightarrow> sem"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   338
where
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   339
  "eval t = evals ENil t"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   340
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   341
definition
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   342
  normalize :: "lam \<Rightarrow> lam"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   343
where
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   344
  "normalize t = reify (eval t)"
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   345
dc6007dd9c30 a little further with NBE
Christian Urban <urbanc@in.tum.de>
parents: 2953
diff changeset
   346
end