author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 18 May 2010 11:46:58 +0200 | |
changeset 2154 | b5c030cfa656 |
parent 1664 | aa999d263b10 |
permissions | -rw-r--r-- |
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 |
1664
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
2 |
imports "../Nominal2_Atoms" "../Nominal2_Eqvt" "../Nominal2_Supp" "../Abs" "../Perm" "../Fv" "../Rsp" |
1459
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") |
1664
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
26 |
[[[], [], [(SOME (@{term rbv5}, false), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [(0,NONE), (2,SOME @{term rbv5})]])] *} |
1459
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 |
|
1664
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
34 |
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_rel_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)) *} |
1459
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
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
56 |
lemma alpha5_reflp: |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
57 |
"y \<approx>5 y \<and> (x \<approx>l x \<and> alpha_rbv5 x x)" |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
58 |
apply (rule rtrm5_rlts.induct) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
59 |
apply (simp_all add: alpha5_inj) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
60 |
apply (rule_tac x="0::perm" in exI) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
61 |
apply (simp add: eqvts alpha_gen fresh_star_def fresh_zero_perm) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
62 |
done |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
63 |
|
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
64 |
lemma alpha5_symp: |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
65 |
"(a \<approx>5 b \<longrightarrow> b \<approx>5 a) \<and> |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
66 |
(x \<approx>l y \<longrightarrow> y \<approx>l x) \<and> |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
67 |
(alpha_rbv5 x y \<longrightarrow> alpha_rbv5 y x)" |
1664
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
68 |
apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
69 |
apply (simp_all add: alpha5_inj) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
70 |
apply (erule conjE)+ |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
71 |
apply (erule exE) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
72 |
apply (rule_tac x="-pi" in exI) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
73 |
apply (rule alpha_gen_sym) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
74 |
apply (simp add: alphas) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
75 |
apply (simp add: alpha5_eqvt) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
76 |
apply (simp add: alphas) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
77 |
apply clarify |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
78 |
apply simp |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
79 |
done |
1474
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
80 |
|
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
81 |
lemma alpha5_transp: |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
82 |
"(a \<approx>5 b \<longrightarrow> (\<forall>c. b \<approx>5 c \<longrightarrow> a \<approx>5 c)) \<and> |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
83 |
(x \<approx>l y \<longrightarrow> (\<forall>z. y \<approx>l z \<longrightarrow> x \<approx>l z)) \<and> |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
84 |
(alpha_rbv5 k l \<longrightarrow> (\<forall>m. alpha_rbv5 l m \<longrightarrow> alpha_rbv5 k m))" |
1664
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
85 |
apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
86 |
apply (rule_tac [!] allI) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
87 |
apply (simp_all add: alpha5_inj) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
88 |
apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
89 |
apply (simp_all add: alpha5_inj) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
90 |
defer |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
91 |
apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
92 |
apply (simp_all add: alpha5_inj) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
93 |
apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
94 |
apply (simp_all add: alpha5_inj) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
95 |
apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
96 |
apply (simp_all add: alpha5_inj) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
97 |
apply (erule conjE)+ |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
98 |
apply (erule exE)+ |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
99 |
apply (rule_tac x="pi + pia" in exI) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
100 |
apply (rule alpha_gen_trans) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
101 |
prefer 6 |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
102 |
apply assumption |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
103 |
apply (simp_all add: alphas alpha5_eqvt) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
104 |
apply (clarify) |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
105 |
apply simp |
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1602
diff
changeset
|
106 |
done |
1474
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
107 |
|
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
108 |
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
|
109 |
"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
|
110 |
"equivp alpha_rlts" |
1474
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
111 |
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
112 |
apply (simp_all only: alpha5_reflp) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
113 |
apply (meson alpha5_symp alpha5_transp) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
114 |
apply (meson alpha5_symp alpha5_transp) |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
115 |
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
|
116 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
117 |
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
|
118 |
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
|
119 |
and |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
120 |
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
|
121 |
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
|
122 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
123 |
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
|
124 |
(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
|
125 |
|> 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
|
126 |
|> 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
|
127 |
|> 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
|
128 |
|> 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
|
129 |
|> 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
|
130 |
|> 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
|
131 |
|> 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
|
132 |
|> 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
|
133 |
|> 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
|
134 |
|> 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
|
135 |
*} |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
136 |
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
|
137 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
138 |
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
|
139 |
"(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
|
140 |
"(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
|
141 |
"(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
|
142 |
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
|
143 |
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
|
144 |
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
|
145 |
done |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
146 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
147 |
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
|
148 |
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
|
149 |
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
|
150 |
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
|
151 |
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
|
152 |
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
|
153 |
done |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
154 |
|
1474
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
155 |
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})]))) *} |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
156 |
print_theorems |
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
157 |
|
1575
2c37f5a8c747
alpha_bn_rsp_pre automatized.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1474
diff
changeset
|
158 |
local_setup {* snd o Local_Theory.note ((@{binding alpha_bn_rsp}, []), prove_alpha_bn_rsp [@{term alpha_rtrm5}, @{term alpha_rlts}] @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts} @{thms alpha5_inj alpha_dis} @{thms alpha5_equivp} @{context} (@{term alpha_rbv5}, 1)) *} |
2c37f5a8c747
alpha_bn_rsp_pre automatized.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1474
diff
changeset
|
159 |
thm alpha_bn_rsp |
1474
8a03753e0e02
Finished all proofs in Term5 and Term5n.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
160 |
|
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
161 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
162 |
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
|
163 |
"(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
|
164 |
"(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
|
165 |
"(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
|
166 |
"(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
|
167 |
"(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
|
168 |
"(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
|
169 |
"(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
|
170 |
"(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
|
171 |
"(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
|
172 |
"(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
|
173 |
"(alpha_rlts ===> alpha_rlts ===> op =) alpha_rbv5 alpha_rbv5" |
1575
2c37f5a8c747
alpha_bn_rsp_pre automatized.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1474
diff
changeset
|
174 |
apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp alpha5_reflp alpha_bn_rsp) |
1459
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
175 |
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
|
176 |
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
|
177 |
done |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
178 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
179 |
lemma |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
180 |
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
|
181 |
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
|
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 |
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
|
184 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
185 |
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
|
186 |
begin |
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 |
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
|
189 |
"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
|
190 |
is |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
191 |
"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
|
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 |
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
|
194 |
"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
|
195 |
is |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
196 |
"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
|
197 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
198 |
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
|
199 |
(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
|
200 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
201 |
end |
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
202 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
203 |
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
|
204 |
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
|
205 |
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
|
206 |
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
|
207 |
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
|
208 |
|
d6d22254aeb7
Added Term5 non-recursive. The bug is there only for the recursive case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
209 |
end |