src/HOL/Nominal/Examples/Contexts.thy
author urbanc
Thu, 20 Dec 2007 01:07:21 +0100
changeset 25722 0a104ddb72d9
parent 25499 7e0ad4838ce9
child 25751 a4e69ce247e0
permissions -rw-r--r--
polishing of some proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
     1
(* $Id$ *)
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     2
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     3
theory Contexts
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     4
imports "../Nominal"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     5
begin
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     6
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     7
text {* 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     8
  
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
     9
  We show here that the Plotkin-style of defining
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    10
  reductions relation based on congruence rules is
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    11
  equivalent to the Felleisen-Hieb-style representation 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    12
  based on contexts. 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    13
  
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    14
  The interesting point is that contexts do not bind 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    15
  anything. On the other hand the operation of replacing 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    16
  a term into a context produces an alpha-equivalent term. 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    17
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    18
*}
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    19
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    20
atom_decl name
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    21
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    22
text {* Terms *}
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    23
nominal_datatype lam = 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    24
    Var "name"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    25
  | App "lam" "lam"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    26
  | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    27
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    28
text {* Contexts - the lambda case in contexts does not bind a name *}
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    29
nominal_datatype ctx = 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    30
    Hole
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    31
  | CAppL "ctx" "lam"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    32
  | CAppR "lam" "ctx" 
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    33
  | CLam "name" "ctx"  ("CLam [_]._" [100,100] 100) 
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    34
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    35
text {* Capture-avoiding substitution and three lemmas *}
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    36
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    37
consts subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [100,100,100] 100)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    38
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    39
nominal_primrec
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    40
  "(Var x)[y::=s] = (if x=y then s else (Var x))"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    41
  "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    42
  "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    43
apply(finite_guess)+
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    44
apply(rule TrueI)+
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    45
apply(simp add: abs_fresh)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    46
apply(fresh_guess)+
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    47
done
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    48
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    49
lemma  subst_eqvt[eqvt]:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    50
  fixes pi::"name prm"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    51
  shows "pi\<bullet>t1[x::=t2] = (pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)]"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    52
by (nominal_induct t1 avoiding: x t2 rule: lam.induct)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    53
   (auto simp add: perm_bij fresh_atm fresh_bij)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    54
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    55
lemma subst_fresh:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    56
  fixes x y::"name"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    57
  and   t t'::"lam"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    58
  shows "y\<sharp>([x].t,t') \<Longrightarrow> y\<sharp>t[x::=t']"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    59
by (nominal_induct t avoiding: x y t' rule: lam.inducts)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    60
   (auto simp add: abs_fresh fresh_prod fresh_atm)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    61
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    62
text {* 
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    63
  Replace is the operation that fills a term into a hole. While
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    64
  contexts themselves are not alpha-equivalence classes, the 
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    65
  filling operation produces an alpha-equivalent lambda-term. 
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    66
*}
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    67
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    68
consts 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    69
  replace :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_<_>" [100,100] 100)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    70
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    71
nominal_primrec
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    72
  "Hole<t> = t"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    73
  "(CAppL E t')<t> = App (E<t>) t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    74
  "(CAppR t' E)<t> = App t' (E<t>)"
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    75
  "(CLam [x].E)<t> = Lam [x].(E<t>)" 
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    76
by (rule TrueI)+
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    77
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    78
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    79
lemma alpha_test: 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    80
  shows "(CLam [x].Hole)<Var x> = (CLam [y].Hole)<Var y>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    81
by (auto simp add: alpha lam.inject calc_atm fresh_atm) 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    82
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    83
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    84
lemma replace_eqvt[eqvt]:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    85
  fixes pi:: "name prm"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    86
  shows "pi\<bullet>(E<e>) = (pi\<bullet>E)<(pi\<bullet>e)>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    87
by (nominal_induct E rule: ctx.inducts) (auto)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    88
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    89
lemma replace_fresh:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    90
  fixes x::"name"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    91
  and   E::"ctx"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    92
  and   t::"lam"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    93
  shows "x\<sharp>(E,t) \<Longrightarrow> x\<sharp>E<t>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    94
by (induct E rule: ctx.weak_induct)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    95
   (auto simp add: fresh_prod abs_fresh)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    96
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    97
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
    98
text {* Composition of two contexts *}
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
    99
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   100
consts 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   101
 ctx_replace :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _" [100,100] 100)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   102
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   103
nominal_primrec
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   104
  "Hole \<circ> E' = E'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   105
  "(CAppL E t') \<circ> E' = CAppL (E \<circ> E') t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   106
  "(CAppR t' E) \<circ> E' = CAppR t' (E \<circ> E')"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   107
  "(CLam [x].E) \<circ> E' = CLam [x].(E \<circ> E')"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   108
by (rule TrueI)+
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   109
  
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   110
lemma ctx_compose:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   111
  shows "E1<E2<t>> = (E1 \<circ> E2)<t>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   112
by (induct E1 rule: ctx.weak_induct) (auto)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   113
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   114
lemma ctx_compose_fresh:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   115
  fixes x::"name"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   116
  and   E1 E2::"ctx"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   117
  shows "x\<sharp>(E1,E2) \<Longrightarrow> x\<sharp>(E1\<circ>E2)"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   118
by (induct E1 rule: ctx.weak_induct)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   119
   (auto simp add: fresh_prod)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   120
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
   121
text {* Beta-reduction via contexts *}
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   122
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   123
inductive
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   124
  ctx_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>x _" [80,80] 80) 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   125
where
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   126
  xbeta[intro]: "x\<sharp>(E,t') \<Longrightarrow> E<App (Lam [x].t) t'> \<longrightarrow>x E<t[x::=t']>" 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   127
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   128
equivariance ctx_red
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   129
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   130
nominal_inductive ctx_red
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   131
  by (simp_all add: replace_fresh subst_fresh abs_fresh)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   132
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
   133
text {* Beta-reduction via congruence rules *}
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   134
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   135
inductive
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   136
  cong_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>o _" [80,80] 80) 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   137
where
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   138
  obeta[intro]: "x\<sharp>t' \<Longrightarrow> App (Lam [x].t) t' \<longrightarrow>o t[x::=t']"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   139
| oapp1[intro]: "t \<longrightarrow>o t' \<Longrightarrow> App t t2 \<longrightarrow>o App t' t2"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   140
| oapp2[intro]: "t \<longrightarrow>o t' \<Longrightarrow> App t2 t \<longrightarrow>o App t2 t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   141
| olam[intro]:  "t \<longrightarrow>o t' \<Longrightarrow> Lam [x].t \<longrightarrow>o Lam [x].t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   142
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   143
equivariance cong_red
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   144
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   145
nominal_inductive cong_red
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   146
  by (simp_all add: subst_fresh abs_fresh)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   147
25722
0a104ddb72d9 polishing of some proofs
urbanc
parents: 25499
diff changeset
   148
text {* The proof that shows both relations are equal *}
25499
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   149
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   150
lemma cong_red_ctx:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   151
  assumes a: "t \<longrightarrow>o t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   152
  shows "E<t> \<longrightarrow>o E<t'>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   153
using a
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   154
by (induct E rule: ctx.weak_induct) (auto)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   155
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   156
lemma ctx_red_ctx:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   157
  assumes a: "t \<longrightarrow>x t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   158
  shows "E<t> \<longrightarrow>x E<t'>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   159
using a 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   160
by (nominal_induct t t' avoiding: E rule: ctx_red.strong_induct) 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   161
   (auto simp add: ctx_compose ctx_compose_fresh)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   162
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   163
lemma ctx_red_hole:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   164
  assumes a: "Hole<t> \<longrightarrow>x Hole<t'>"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   165
  shows "t \<longrightarrow>x t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   166
using a by simp
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   167
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   168
theorem ctx_red_cong_red:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   169
  assumes a: "t \<longrightarrow>x t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   170
  shows "t \<longrightarrow>o t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   171
using a 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   172
by (induct) (auto intro!: cong_red_ctx)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   173
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   174
theorem cong_red_ctx_red:
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   175
  assumes a: "t \<longrightarrow>o t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   176
  shows "t \<longrightarrow>x t'"
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   177
using a 
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   178
apply(induct)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   179
apply(rule ctx_red_hole)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   180
apply(rule xbeta)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   181
apply(simp)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   182
apply(metis ctx_red_ctx replace.simps)+ (* found by SledgeHammer *)
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   183
done
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   184
7e0ad4838ce9 an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff changeset
   185
end