985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory LFex
|
1219
|
2 |
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
begin
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
atom_decl ident
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
|
1234
|
8 |
datatype rkind =
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
Type
|
1234
|
10 |
| KPi "rty" "name" "rkind"
|
|
11 |
and rty =
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
TConst "ident"
|
1234
|
13 |
| TApp "rty" "rtrm"
|
|
14 |
| TPi "rty" "name" "rty"
|
|
15 |
and rtrm =
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
Const "ident"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
17 |
| Var "name"
|
1234
|
18 |
| App "rtrm" "rtrm"
|
|
19 |
| Lam "rty" "name" "rtrm"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
|
1234
|
21 |
|
1239
|
22 |
setup {* snd o define_raw_perms ["rkind", "rty", "rtrm"] ["LFex.rkind", "LFex.rty", "LFex.rtrm"] *}
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
|
1219
|
24 |
local_setup {*
|
1234
|
25 |
snd o define_fv_alpha "LFex.rkind"
|
1219
|
26 |
[[ [], [[], [(NONE, 1)], [(NONE, 1)]] ],
|
|
27 |
[ [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]] ],
|
|
28 |
[ [[]], [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]]]] *}
|
|
29 |
notation
|
1234
|
30 |
alpha_rkind ("_ \<approx>ki _" [100, 100] 100)
|
1237
|
31 |
and alpha_rty ("_ \<approx>ty _" [100, 100] 100)
|
1234
|
32 |
and alpha_rtrm ("_ \<approx>tr _" [100, 100] 100)
|
|
33 |
thm fv_rkind_fv_rty_fv_rtrm.simps alpha_rkind_alpha_rty_alpha_rtrm.intros
|
1238
|
34 |
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_rkind_alpha_rty_alpha_rtrm_inj}, []), (build_alpha_inj @{thms alpha_rkind_alpha_rty_alpha_rtrm.intros} @{thms rkind.distinct rty.distinct rtrm.distinct rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} ctxt)) ctxt)) *}
|
|
35 |
thm alpha_rkind_alpha_rty_alpha_rtrm_inj
|
1219
|
36 |
|
992
|
37 |
lemma rfv_eqvt[eqvt]:
|
1236
|
38 |
"((pi\<bullet>fv_rkind t1) = fv_rkind (pi\<bullet>t1))"
|
|
39 |
"((pi\<bullet>fv_rty t2) = fv_rty (pi\<bullet>t2))"
|
|
40 |
"((pi\<bullet>fv_rtrm t3) = fv_rtrm (pi\<bullet>t3))"
|
1234
|
41 |
apply(induct t1 and t2 and t3 rule: rkind_rty_rtrm.inducts)
|
1239
|
42 |
apply(simp_all add: union_eqvt Diff_eqvt)
|
993
|
43 |
apply(simp_all add: permute_set_eq atom_eqvt)
|
992
|
44 |
done
|
|
45 |
|
|
46 |
lemma alpha_eqvt:
|
|
47 |
"t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)"
|
1237
|
48 |
"t2 \<approx>ty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>ty (pi \<bullet> s2)"
|
992
|
49 |
"t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)"
|
1237
|
50 |
apply(induct rule: alpha_rkind_alpha_rty_alpha_rtrm.inducts)
|
|
51 |
apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm.intros)
|
|
52 |
apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm_inj)
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
53 |
apply (rule alpha_gen_atom_eqvt)
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
54 |
apply (simp add: rfv_eqvt)
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
55 |
apply assumption
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
56 |
apply (rule alpha_gen_atom_eqvt)
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
57 |
apply (simp add: rfv_eqvt)
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
58 |
apply assumption
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
59 |
apply (rule alpha_gen_atom_eqvt)
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
60 |
apply (simp add: rfv_eqvt)
|
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
61 |
apply assumption
|
992
|
62 |
done
|
|
63 |
|
1239
|
64 |
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_equivps}, []),
|
|
65 |
(build_equivps [@{term alpha_rkind}, @{term alpha_rty}, @{term alpha_rtrm}]
|
|
66 |
@{thm rkind_rty_rtrm.induct} @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct}
|
|
67 |
@{thms rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj}
|
|
68 |
@{thms rkind.distinct rty.distinct rtrm.distinct}
|
|
69 |
@{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases}
|
|
70 |
@{thms alpha_eqvt} ctxt)) ctxt)) *}
|
|
71 |
thm alpha_equivps
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
72 |
|
1237
|
73 |
quotient_type RKIND = rkind / alpha_rkind
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
74 |
by (rule alpha_equivps)
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
75 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
76 |
quotient_type
|
1237
|
77 |
RTY = rty / alpha_rty and
|
|
78 |
RTRM = rtrm / alpha_rtrm
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
by (auto intro: alpha_equivps)
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
80 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
81 |
quotient_definition
|
1234
|
82 |
"TYP :: RKIND"
|
1139
|
83 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
84 |
"Type"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
85 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
86 |
quotient_definition
|
1234
|
87 |
"KPI :: RTY \<Rightarrow> name \<Rightarrow> RKIND \<Rightarrow> RKIND"
|
1139
|
88 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
89 |
"KPi"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
90 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
91 |
quotient_definition
|
1234
|
92 |
"TCONST :: ident \<Rightarrow> RTY"
|
1139
|
93 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
94 |
"TConst"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
95 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
96 |
quotient_definition
|
1234
|
97 |
"TAPP :: RTY \<Rightarrow> RTRM \<Rightarrow> RTY"
|
1139
|
98 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
99 |
"TApp"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
100 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
101 |
quotient_definition
|
1234
|
102 |
"TPI :: RTY \<Rightarrow> name \<Rightarrow> RTY \<Rightarrow> RTY"
|
1139
|
103 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
104 |
"TPi"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
105 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
106 |
(* FIXME: does not work with CONST *)
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
107 |
quotient_definition
|
1234
|
108 |
"CONS :: ident \<Rightarrow> RTRM"
|
1139
|
109 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
110 |
"Const"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
111 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
112 |
quotient_definition
|
1234
|
113 |
"VAR :: name \<Rightarrow> RTRM"
|
1139
|
114 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
115 |
"Var"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
116 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
117 |
quotient_definition
|
1234
|
118 |
"APP :: RTRM \<Rightarrow> RTRM \<Rightarrow> RTRM"
|
1139
|
119 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
120 |
"App"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
121 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
122 |
quotient_definition
|
1234
|
123 |
"LAM :: RTY \<Rightarrow> name \<Rightarrow> RTRM \<Rightarrow> RTRM"
|
1139
|
124 |
is
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
125 |
"Lam"
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
126 |
|
1234
|
127 |
(* FIXME: print out a warning if the type contains a liftet type, like rkind \<Rightarrow> name set *)
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
128 |
quotient_definition
|
1236
|
129 |
"fv_kind :: RKIND \<Rightarrow> atom set"
|
1139
|
130 |
is
|
1236
|
131 |
"fv_rkind"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
132 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
133 |
quotient_definition
|
1236
|
134 |
"fv_ty :: RTY \<Rightarrow> atom set"
|
1139
|
135 |
is
|
1236
|
136 |
"fv_rty"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
137 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
138 |
quotient_definition
|
1236
|
139 |
"fv_trm :: RTRM \<Rightarrow> atom set"
|
1139
|
140 |
is
|
1236
|
141 |
"fv_rtrm"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
142 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
143 |
lemma alpha_rfv:
|
1236
|
144 |
shows "(t \<approx>ki s \<longrightarrow> fv_rkind t = fv_rkind s) \<and>
|
1237
|
145 |
(t1 \<approx>ty s1 \<longrightarrow> fv_rty t1 = fv_rty s1) \<and>
|
1236
|
146 |
(t2 \<approx>tr s2 \<longrightarrow> fv_rtrm t2 = fv_rtrm s2)"
|
1237
|
147 |
apply(rule alpha_rkind_alpha_rty_alpha_rtrm.induct)
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
148 |
apply(simp_all add: alpha_gen)
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
149 |
done
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
150 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
151 |
lemma perm_rsp[quot_respect]:
|
1237
|
152 |
"(op = ===> alpha_rkind ===> alpha_rkind) permute permute"
|
|
153 |
"(op = ===> alpha_rty ===> alpha_rty) permute permute"
|
|
154 |
"(op = ===> alpha_rtrm ===> alpha_rtrm) permute permute"
|
992
|
155 |
by (simp_all add:alpha_eqvt)
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
156 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
157 |
lemma tconst_rsp[quot_respect]:
|
1237
|
158 |
"(op = ===> alpha_rty) TConst TConst"
|
|
159 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros) done
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
160 |
lemma tapp_rsp[quot_respect]:
|
1237
|
161 |
"(alpha_rty ===> alpha_rtrm ===> alpha_rty) TApp TApp"
|
|
162 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros) done
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
163 |
lemma var_rsp[quot_respect]:
|
1237
|
164 |
"(op = ===> alpha_rtrm) Var Var"
|
|
165 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros) done
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
166 |
lemma app_rsp[quot_respect]:
|
1237
|
167 |
"(alpha_rtrm ===> alpha_rtrm ===> alpha_rtrm) App App"
|
|
168 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros) done
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
169 |
lemma const_rsp[quot_respect]:
|
1237
|
170 |
"(op = ===> alpha_rtrm) Const Const"
|
|
171 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros) done
|
992
|
172 |
|
|
173 |
lemma kpi_rsp[quot_respect]:
|
1237
|
174 |
"(alpha_rty ===> op = ===> alpha_rkind ===> alpha_rkind) KPi KPi"
|
|
175 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros)
|
|
176 |
apply (rule alpha_rkind_alpha_rty_alpha_rtrm.intros(2)) apply simp_all
|
993
|
177 |
apply (rule_tac x="0" in exI)
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
178 |
apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen)
|
993
|
179 |
done
|
|
180 |
|
992
|
181 |
lemma tpi_rsp[quot_respect]:
|
1237
|
182 |
"(alpha_rty ===> op = ===> alpha_rty ===> alpha_rty) TPi TPi"
|
|
183 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros)
|
|
184 |
apply (rule alpha_rkind_alpha_rty_alpha_rtrm.intros(5)) apply simp_all
|
993
|
185 |
apply (rule_tac x="0" in exI)
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
186 |
apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen)
|
993
|
187 |
done
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
188 |
lemma lam_rsp[quot_respect]:
|
1237
|
189 |
"(alpha_rty ===> op = ===> alpha_rtrm ===> alpha_rtrm) Lam Lam"
|
|
190 |
apply (auto intro: alpha_rkind_alpha_rty_alpha_rtrm.intros)
|
|
191 |
apply (rule alpha_rkind_alpha_rty_alpha_rtrm.intros(9)) apply simp_all
|
993
|
192 |
apply (rule_tac x="0" in exI)
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
193 |
apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen)
|
993
|
194 |
done
|
992
|
195 |
|
1236
|
196 |
lemma fv_rty_rsp[quot_respect]:
|
1237
|
197 |
"(alpha_rty ===> op =) fv_rty fv_rty"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
198 |
by (simp add: alpha_rfv)
|
1236
|
199 |
lemma fv_rkind_rsp[quot_respect]:
|
1237
|
200 |
"(alpha_rkind ===> op =) fv_rkind fv_rkind"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
201 |
by (simp add: alpha_rfv)
|
1236
|
202 |
lemma fv_rtrm_rsp[quot_respect]:
|
1237
|
203 |
"(alpha_rtrm ===> op =) fv_rtrm fv_rtrm"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
204 |
by (simp add: alpha_rfv)
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
205 |
|
1234
|
206 |
thm rkind_rty_rtrm.induct
|
|
207 |
lemmas RKIND_RTY_RTRM_induct = rkind_rty_rtrm.induct[quot_lifted]
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
208 |
|
1234
|
209 |
thm rkind_rty_rtrm.inducts
|
|
210 |
lemmas RKIND_RTY_RTRM_inducts = rkind_rty_rtrm.inducts[quot_lifted]
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
211 |
|
1234
|
212 |
instantiation RKIND and RTY and RTRM :: pt
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
213 |
begin
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
214 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
215 |
quotient_definition
|
1234
|
216 |
"permute_RKIND :: perm \<Rightarrow> RKIND \<Rightarrow> RKIND"
|
1139
|
217 |
is
|
1234
|
218 |
"permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
219 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
220 |
quotient_definition
|
1234
|
221 |
"permute_RTY :: perm \<Rightarrow> RTY \<Rightarrow> RTY"
|
1139
|
222 |
is
|
1234
|
223 |
"permute :: perm \<Rightarrow> rty \<Rightarrow> rty"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
224 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
225 |
quotient_definition
|
1234
|
226 |
"permute_RTRM :: perm \<Rightarrow> RTRM \<Rightarrow> RTRM"
|
1139
|
227 |
is
|
1234
|
228 |
"permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
229 |
|
1234
|
230 |
lemmas permute_ktt[simp] = permute_rkind_permute_rty_permute_rtrm.simps[quot_lifted]
|
1083
|
231 |
|
1234
|
232 |
lemma perm_zero_ok: "0 \<bullet> (x :: RKIND) = x \<and> 0 \<bullet> (y :: RTY) = y \<and> 0 \<bullet> (z :: RTRM) = z"
|
|
233 |
apply (induct rule: RKIND_RTY_RTRM_induct)
|
1083
|
234 |
apply (simp_all)
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
235 |
done
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
236 |
|
1082
|
237 |
lemma perm_add_ok:
|
1234
|
238 |
"((p + q) \<bullet> (x1 :: RKIND) = (p \<bullet> q \<bullet> x1))"
|
|
239 |
"((p + q) \<bullet> (x2 :: RTY) = p \<bullet> q \<bullet> x2)"
|
|
240 |
"((p + q) \<bullet> (x3 :: RTRM) = p \<bullet> q \<bullet> x3)"
|
|
241 |
apply (induct x1 and x2 and x3 rule: RKIND_RTY_RTRM_inducts)
|
1083
|
242 |
apply (simp_all)
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
243 |
done
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
244 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
245 |
instance
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
246 |
apply default
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
247 |
apply (simp_all add: perm_zero_ok perm_add_ok)
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
248 |
done
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
249 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
250 |
end
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
251 |
|
1237
|
252 |
lemmas ALPHA_RKIND_ALPHA_RTY_ALPHA_RTRM_inducts = alpha_rkind_alpha_rty_alpha_rtrm.inducts[unfolded alpha_gen, quot_lifted, folded alpha_gen]
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
253 |
|
1237
|
254 |
lemmas RKIND_RTY_RTRM_INJECT = alpha_rkind_alpha_rty_alpha_rtrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
255 |
|
1236
|
256 |
lemmas fv_kind_ty_trm = fv_rkind_fv_rty_fv_rtrm.simps[quot_lifted]
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
257 |
|
1082
|
258 |
lemmas fv_eqvt = rfv_eqvt[quot_lifted]
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
259 |
|
1234
|
260 |
lemma supp_rkind_rty_rtrm_easy:
|
1002
|
261 |
"supp TYP = {}"
|
|
262 |
"supp (TCONST i) = {atom i}"
|
|
263 |
"supp (TAPP A M) = supp A \<union> supp M"
|
|
264 |
"supp (CONS i) = {atom i}"
|
|
265 |
"supp (VAR x) = {atom x}"
|
|
266 |
"supp (APP M N) = supp M \<union> supp N"
|
1239
|
267 |
apply (simp_all add: supp_def permute_ktt)
|
|
268 |
apply (simp_all only: RKIND_RTY_RTRM_INJECT)
|
1002
|
269 |
apply (simp_all only: supp_at_base[simplified supp_def])
|
|
270 |
apply (simp_all add: Collect_imp_eq Collect_neg_eq)
|
|
271 |
done
|
|
272 |
|
|
273 |
lemma supp_bind:
|
|
274 |
"(supp (atom na, (ty, ki))) supports (KPI ty na ki)"
|
|
275 |
"(supp (atom na, (ty, ty2))) supports (TPI ty na ty2)"
|
1234
|
276 |
"(supp (atom na, (ty, rtrm))) supports (LAM ty na rtrm)"
|
1002
|
277 |
apply(simp_all add: supports_def)
|
|
278 |
apply(fold fresh_def)
|
|
279 |
apply(simp_all add: fresh_Pair swap_fresh_fresh)
|
|
280 |
apply(clarify)
|
|
281 |
apply(subst swap_at_base_simps(3))
|
|
282 |
apply(simp_all add: fresh_atom)
|
|
283 |
apply(clarify)
|
|
284 |
apply(subst swap_at_base_simps(3))
|
|
285 |
apply(simp_all add: fresh_atom)
|
|
286 |
apply(clarify)
|
|
287 |
apply(subst swap_at_base_simps(3))
|
|
288 |
apply(simp_all add: fresh_atom)
|
|
289 |
done
|
|
290 |
|
1234
|
291 |
lemma RKIND_RTY_RTRM_fs:
|
|
292 |
"finite (supp (x\<Colon>RKIND))"
|
|
293 |
"finite (supp (y\<Colon>RTY))"
|
|
294 |
"finite (supp (z\<Colon>RTRM))"
|
|
295 |
apply(induct x and y and z rule: RKIND_RTY_RTRM_inducts)
|
|
296 |
apply(simp_all add: supp_rkind_rty_rtrm_easy)
|
1002
|
297 |
apply(rule supports_finite)
|
|
298 |
apply(rule supp_bind(1))
|
|
299 |
apply(simp add: supp_Pair supp_atom)
|
|
300 |
apply(rule supports_finite)
|
|
301 |
apply(rule supp_bind(2))
|
|
302 |
apply(simp add: supp_Pair supp_atom)
|
|
303 |
apply(rule supports_finite)
|
|
304 |
apply(rule supp_bind(3))
|
|
305 |
apply(simp add: supp_Pair supp_atom)
|
|
306 |
done
|
|
307 |
|
1234
|
308 |
instance RKIND and RTY and RTRM :: fs
|
1002
|
309 |
apply(default)
|
1234
|
310 |
apply(simp_all only: RKIND_RTY_RTRM_fs)
|
1002
|
311 |
done
|
|
312 |
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
313 |
lemma supp_fv:
|
1236
|
314 |
"supp t1 = fv_kind t1"
|
|
315 |
"supp t2 = fv_ty t2"
|
|
316 |
"supp t3 = fv_trm t3"
|
1234
|
317 |
apply(induct t1 and t2 and t3 rule: RKIND_RTY_RTRM_inducts)
|
|
318 |
apply (simp_all add: supp_rkind_rty_rtrm_easy)
|
1236
|
319 |
apply (simp_all add: fv_kind_ty_trm)
|
1234
|
320 |
apply(subgoal_tac "supp (KPI rty name rkind) = supp rty \<union> supp (Abs {atom name} rkind)")
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
321 |
apply(simp add: supp_Abs Set.Un_commute)
|
1002
|
322 |
apply(simp (no_asm) add: supp_def)
|
1234
|
323 |
apply(simp add: RKIND_RTY_RTRM_INJECT)
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
324 |
apply(simp add: Abs_eq_iff)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
325 |
apply(simp add: alpha_gen)
|
1045
|
326 |
apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute insert_eqvt empty_eqvt atom_eqvt)
|
|
327 |
apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric])
|
1234
|
328 |
apply(subgoal_tac "supp (TPI rty1 name rty2) = supp rty1 \<union> supp (Abs {atom name} rty2)")
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
329 |
apply(simp add: supp_Abs Set.Un_commute)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
330 |
apply(simp (no_asm) add: supp_def)
|
1234
|
331 |
apply(simp add: RKIND_RTY_RTRM_INJECT)
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
332 |
apply(simp add: Abs_eq_iff)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
333 |
apply(simp add: alpha_gen)
|
1045
|
334 |
apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric] insert_eqvt empty_eqvt atom_eqvt)
|
1002
|
335 |
apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute)
|
1234
|
336 |
apply(subgoal_tac "supp (LAM rty name rtrm) = supp rty \<union> supp (Abs {atom name} rtrm)")
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
337 |
apply(simp add: supp_Abs Set.Un_commute)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
338 |
apply(simp (no_asm) add: supp_def)
|
1234
|
339 |
apply(simp add: RKIND_RTY_RTRM_INJECT)
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
340 |
apply(simp add: Abs_eq_iff)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
341 |
apply(simp add: alpha_gen)
|
1045
|
342 |
apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric] insert_eqvt empty_eqvt atom_eqvt)
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
343 |
apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
344 |
done
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
345 |
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
346 |
(* Not needed anymore *)
|
1027
163d6917af62
LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
347 |
lemma supp_kpi_pre: "supp (KPI A x K) = (supp (Abs {atom x} K)) \<union> supp A"
|
1234
|
348 |
apply (simp add: permute_set_eq supp_def Abs_eq_iff RKIND_RTY_RTRM_INJECT)
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
349 |
apply (simp add: alpha_gen supp_fv)
|
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
350 |
apply (simp add: Collect_imp_eq Collect_neg_eq add: atom_eqvt Set.Un_commute)
|
997
b7d259ded92e
Ported LF to the generic lambda and solved the simpler _supp cases.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
351 |
done
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
352 |
|
1234
|
353 |
lemma supp_rkind_rty_rtrm:
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
354 |
"supp TYP = {}"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
355 |
"supp (KPI A x K) = supp A \<union> (supp K - {atom x})"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
356 |
"supp (TCONST i) = {atom i}"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
357 |
"supp (TAPP A M) = supp A \<union> supp M"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
358 |
"supp (TPI A x B) = supp A \<union> (supp B - {atom x})"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
359 |
"supp (CONS i) = {atom i}"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
360 |
"supp (VAR x) = {atom x}"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
361 |
"supp (APP M N) = supp M \<union> supp N"
|
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
362 |
"supp (LAM A x M) = supp A \<union> (supp M - {atom x})"
|
1234
|
363 |
apply (simp_all only: supp_rkind_rty_rtrm_easy)
|
1236
|
364 |
apply (simp_all only: supp_fv fv_kind_ty_trm)
|
1036
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
365 |
done
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
366 |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
367 |
end
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
368 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
369 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
370 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
371 |
|