1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 1
theory Term5
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 3
begin
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 4
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 5
atom_decl name
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 6
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 7
datatype rtrm5 =
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 8
rVr5 "name"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 9
| rAp5 "rtrm5" "rtrm5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 10
| rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 11
and rlts =
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 12
rLnil
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 13
| rLcons "name" "rtrm5" "rlts"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 14
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 15
primrec
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 16
rbv5
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 17
where
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 18
"rbv5 rLnil = {}"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 19
| "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 20
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 21
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 22
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term5.rtrm5") 2 *}
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 23
print_theorems
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 24
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 25
local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "Term5.rtrm5")
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 26
[[[], [], [(SOME (@{term rbv5}, false), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [0, 2]])] *}
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 27
print_theorems
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 28
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 29
notation
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 30
alpha_rtrm5 ("_ \<approx>5 _" [100, 100] 100) and
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 31
alpha_rlts ("_ \<approx>l _" [100, 100] 100)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 32
thm alpha_rtrm5_alpha_rlts_alpha_rbv5.intros
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 33
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 34
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_alpha_inj @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} ctxt)) ctxt)) *}
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 35
thm alpha5_inj
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 36
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 37
lemma rbv5_eqvt[eqvt]:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 38
"pi \<bullet> (rbv5 x) = rbv5 (pi \<bullet> x)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 39
apply (induct x)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 40
apply (simp_all add: eqvts atom_eqvt)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 41
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 42
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 43
lemma fv_rtrm5_rlts_eqvt[eqvt]:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 44
"pi \<bullet> (fv_rtrm5 x) = fv_rtrm5 (pi \<bullet> x)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 45
"pi \<bullet> (fv_rlts l) = fv_rlts (pi \<bullet> l)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 46
"pi \<bullet> (fv_rbv5 l) = fv_rbv5 (pi \<bullet> l)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 47
apply (induct x and l)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 48
apply (simp_all add: eqvts atom_eqvt)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 49
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 50
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 51
local_setup {*
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 52
(fn ctxt => snd (Local_Theory.note ((@{binding alpha5_eqvt}, []),
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 53
build_alpha_eqvts [@{term alpha_rtrm5}, @{term alpha_rlts}, @{term alpha_rbv5}] (fn _ => alpha_eqvt_tac @{thm alpha_rtrm5_alpha_rlts_alpha_rbv5.induct} @{thms alpha5_inj permute_rtrm5_permute_rlts.simps} ctxt 1) ctxt) ctxt)) *}
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 54
print_theorems
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 55
1474
+ − 56
lemma alpha5_reflp:
+ − 57
"y \<approx>5 y \<and> (x \<approx>l x \<and> alpha_rbv5 x x)"
+ − 58
apply (rule rtrm5_rlts.induct)
+ − 59
apply (simp_all add: alpha5_inj)
+ − 60
apply (rule_tac x="0::perm" in exI)
+ − 61
apply (simp add: eqvts alpha_gen fresh_star_def fresh_zero_perm)
+ − 62
done
+ − 63
+ − 64
lemma alpha5_symp:
+ − 65
"(a \<approx>5 b \<longrightarrow> b \<approx>5 a) \<and>
+ − 66
(x \<approx>l y \<longrightarrow> y \<approx>l x) \<and>
+ − 67
(alpha_rbv5 x y \<longrightarrow> alpha_rbv5 y x)"
+ − 68
sorry
+ − 69
+ − 70
lemma alpha5_transp:
+ − 71
"(a \<approx>5 b \<longrightarrow> (\<forall>c. b \<approx>5 c \<longrightarrow> a \<approx>5 c)) \<and>
+ − 72
(x \<approx>l y \<longrightarrow> (\<forall>z. y \<approx>l z \<longrightarrow> x \<approx>l z)) \<and>
+ − 73
(alpha_rbv5 k l \<longrightarrow> (\<forall>m. alpha_rbv5 l m \<longrightarrow> alpha_rbv5 k m))"
+ − 74
sorry
+ − 75
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 76
lemma alpha5_equivp:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 77
"equivp alpha_rtrm5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 78
"equivp alpha_rlts"
1474
+ − 79
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
+ − 80
apply (simp_all only: alpha5_reflp)
+ − 81
apply (meson alpha5_symp alpha5_transp)
+ − 82
apply (meson alpha5_symp alpha5_transp)
+ − 83
done
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 84
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 85
quotient_type
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 86
trm5 = rtrm5 / alpha_rtrm5
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 87
and
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 88
lts = rlts / alpha_rlts
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 89
by (auto intro: alpha5_equivp)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 90
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 91
local_setup {*
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 92
(fn ctxt => ctxt
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 93
|> snd o (Quotient_Def.quotient_lift_const ("Vr5", @{term rVr5}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 94
|> snd o (Quotient_Def.quotient_lift_const ("Ap5", @{term rAp5}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 95
|> snd o (Quotient_Def.quotient_lift_const ("Lt5", @{term rLt5}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 96
|> snd o (Quotient_Def.quotient_lift_const ("Lnil", @{term rLnil}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 97
|> snd o (Quotient_Def.quotient_lift_const ("Lcons", @{term rLcons}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 98
|> snd o (Quotient_Def.quotient_lift_const ("fv_trm5", @{term fv_rtrm5}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 99
|> snd o (Quotient_Def.quotient_lift_const ("fv_lts", @{term fv_rlts}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 100
|> snd o (Quotient_Def.quotient_lift_const ("fv_bv5", @{term fv_rbv5}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 101
|> snd o (Quotient_Def.quotient_lift_const ("bv5", @{term rbv5}))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 102
|> snd o (Quotient_Def.quotient_lift_const ("alpha_bv5", @{term alpha_rbv5})))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 103
*}
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 104
print_theorems
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 105
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 106
lemma alpha5_rfv:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 107
"(t \<approx>5 s \<Longrightarrow> fv_rtrm5 t = fv_rtrm5 s)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 108
"(l \<approx>l m \<Longrightarrow> (fv_rlts l = fv_rlts m \<and> fv_rbv5 l = fv_rbv5 m))"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 109
"(alpha_rbv5 b c \<Longrightarrow> fv_rbv5 b = fv_rbv5 c)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 110
apply(induct rule: alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 111
apply(simp_all)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 112
apply(simp add: alpha_gen)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 113
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 114
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 115
lemma bv_list_rsp:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 116
shows "x \<approx>l y \<Longrightarrow> rbv5 x = rbv5 y"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 117
apply(induct rule: alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts(2))
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 118
apply(simp_all)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 119
apply(clarify)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 120
apply simp
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 121
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 122
1474
+ − 123
local_setup {* snd o Local_Theory.note ((@{binding alpha_dis}, []), (flat (map (distinct_rel @{context} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases}) [(@{thms rtrm5.distinct}, @{term alpha_rtrm5}), (@{thms rlts.distinct}, @{term alpha_rlts}), (@{thms rlts.distinct}, @{term alpha_rbv5})]))) *}
+ − 124
print_theorems
+ − 125
+ − 126
lemma alpha_rbv_rsp_pre:
+ − 127
"x \<approx>l y \<Longrightarrow> \<forall>z. alpha_rbv5 x z = alpha_rbv5 y z"
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 128
apply (erule alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts(2))
1474
+ − 129
apply (simp_all add: alpha_dis alpha5_inj)
+ − 130
apply clarify
+ − 131
apply (case_tac [!] z)
+ − 132
apply (simp_all add: alpha_dis alpha5_inj)
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 133
apply clarify
1474
+ − 134
apply auto
+ − 135
apply (meson alpha5_symp alpha5_transp)
+ − 136
apply (meson alpha5_symp alpha5_transp)
+ − 137
done
+ − 138
+ − 139
lemma alpha_rbv_rsp_pre2:
+ − 140
"x \<approx>l y \<Longrightarrow> \<forall>z. alpha_rbv5 z x = alpha_rbv5 z y"
+ − 141
apply (erule alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts(2))
+ − 142
apply (simp_all add: alpha_dis alpha5_inj)
+ − 143
apply clarify
+ − 144
apply (case_tac [!] z)
+ − 145
apply (simp_all add: alpha_dis alpha5_inj)
+ − 146
apply clarify
+ − 147
apply auto
+ − 148
apply (meson alpha5_symp alpha5_transp)
+ − 149
apply (meson alpha5_symp alpha5_transp)
+ − 150
done
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 151
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 152
lemma [quot_respect]:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 153
"(alpha_rlts ===> op =) fv_rlts fv_rlts"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 154
"(alpha_rlts ===> op =) fv_rbv5 fv_rbv5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 155
"(alpha_rtrm5 ===> op =) fv_rtrm5 fv_rtrm5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 156
"(alpha_rlts ===> op =) rbv5 rbv5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 157
"(op = ===> alpha_rtrm5) rVr5 rVr5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 158
"(alpha_rtrm5 ===> alpha_rtrm5 ===> alpha_rtrm5) rAp5 rAp5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 159
"(alpha_rlts ===> alpha_rtrm5 ===> alpha_rtrm5) rLt5 rLt5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 160
"(op = ===> alpha_rtrm5 ===> alpha_rlts ===> alpha_rlts) rLcons rLcons"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 161
"(op = ===> alpha_rtrm5 ===> alpha_rtrm5) permute permute"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 162
"(op = ===> alpha_rlts ===> alpha_rlts) permute permute"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 163
"(alpha_rlts ===> alpha_rlts ===> op =) alpha_rbv5 alpha_rbv5"
1474
+ − 164
apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp alpha_rbv_rsp_pre alpha_rbv_rsp_pre2 alpha5_reflp)
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 165
apply (clarify)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 166
apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 167
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 168
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 169
lemma
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 170
shows "(alpha_rlts ===> op =) rbv5 rbv5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 171
by (simp add: bv_list_rsp)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 172
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 173
lemmas trm5_lts_inducts = rtrm5_rlts.inducts[quot_lifted]
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 174
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 175
instantiation trm5 and lts :: pt
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 176
begin
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 177
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 178
quotient_definition
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 179
"permute_trm5 :: perm \<Rightarrow> trm5 \<Rightarrow> trm5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 180
is
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 181
"permute :: perm \<Rightarrow> rtrm5 \<Rightarrow> rtrm5"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 182
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 183
quotient_definition
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 184
"permute_lts :: perm \<Rightarrow> lts \<Rightarrow> lts"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 185
is
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 186
"permute :: perm \<Rightarrow> rlts \<Rightarrow> rlts"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 187
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 188
instance by default
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 189
(simp_all add: permute_rtrm5_permute_rlts_zero[quot_lifted] permute_rtrm5_permute_rlts_append[quot_lifted])
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 190
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 191
end
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 192
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 193
lemmas permute_trm5_lts = permute_rtrm5_permute_rlts.simps[quot_lifted]
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 194
lemmas bv5[simp] = rbv5.simps[quot_lifted]
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 195
lemmas fv_trm5_bv5[simp] = fv_rtrm5_fv_rbv5.simps[quot_lifted]
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 196
lemmas fv_lts[simp] = fv_rlts.simps[quot_lifted]
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 197
lemmas alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 198
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 199
lemma lets_bla:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 200
"x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Lt5 (Lcons x (Vr5 y) Lnil) (Vr5 x)) \<noteq> (Lt5 (Lcons x (Vr5 z) Lnil) (Vr5 x))"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 201
apply (simp only: alpha5_INJ)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 202
apply (simp only: bv5)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 203
apply simp
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 204
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 205
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 206
lemma lets_ok:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 207
"(Lt5 (Lcons x (Vr5 y) Lnil) (Vr5 x)) = (Lt5 (Lcons y (Vr5 y) Lnil) (Vr5 y))"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 208
apply (simp add: alpha5_INJ)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 209
apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 210
apply (simp_all add: alpha_gen)
1474
+ − 211
apply (simp add: permute_trm5_lts fresh_star_def eqvts)
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 212
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 213
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 214
lemma lets_ok3:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 215
"x \<noteq> y \<Longrightarrow>
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 216
(Lt5 (Lcons x (Ap5 (Vr5 y) (Vr5 x)) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 217
(Lt5 (Lcons y (Ap5 (Vr5 x) (Vr5 y)) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 218
apply (simp add: permute_trm5_lts alpha_gen alpha5_INJ)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 219
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 220
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 221
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 222
lemma lets_not_ok1:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 223
"(Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) =
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 224
(Lt5 (Lcons y (Vr5 x) (Lcons x (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 225
apply (simp add: alpha5_INJ alpha_gen)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 226
apply (rule_tac x="0::perm" in exI)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 227
apply (simp add: permute_trm5_lts fresh_star_def alpha5_INJ(5) alpha5_INJ(2) alpha5_INJ(1) eqvts)
1474
+ − 228
apply blast
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 229
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 230
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 231
lemma distinct_helper:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 232
shows "\<not>(rVr5 x \<approx>5 rAp5 y z)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 233
apply auto
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 234
apply (erule alpha_rtrm5.cases)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 235
apply (simp_all only: rtrm5.distinct)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 236
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 237
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 238
lemma distinct_helper2:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 239
shows "(Vr5 x) \<noteq> (Ap5 y z)"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 240
by (lifting distinct_helper)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 241
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 242
lemma lets_nok:
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 243
"x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow>
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 244
(Lt5 (Lcons x (Ap5 (Vr5 z) (Vr5 z)) (Lcons y (Vr5 z) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 245
(Lt5 (Lcons y (Vr5 z) (Lcons x (Ap5 (Vr5 z) (Vr5 z)) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 246
apply (simp only: alpha5_INJ(3) alpha5_INJ(5) alpha_gen permute_trm5_lts fresh_star_def)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 247
apply (simp add: distinct_helper2 alpha5_INJ permute_trm5_lts)
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 248
done
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 249
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
+ − 250
end