author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 18 May 2010 17:56:41 +0200 | |
changeset 2160 | 8c24ee88b8e8 |
parent 2120 | 2786ff1df475 |
permissions | -rw-r--r-- |
1600 | 1 |
theory ExLet |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
2 |
imports "../NewParser" "../../Attic/Prove" |
1600 | 3 |
begin |
4 |
||
5 |
text {* example 3 or example 5 from Terms.thy *} |
|
6 |
||
7 |
atom_decl name |
|
8 |
||
9 |
nominal_datatype trm = |
|
10 |
Vr "name" |
|
11 |
| Ap "trm" "trm" |
|
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
12 |
| Lm x::"name" t::"trm" bind_set x in t |
1602
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
13 |
| Lt a::"lts" t::"trm" bind "bn a" in t |
1731
3832a31a73b1
a test with let having multiple bodies
Christian Urban <urbanc@in.tum.de>
parents:
1685
diff
changeset
|
14 |
(*| L a::"lts" t1::"trm" t2::"trm" bind "bn a" in t1, bind "bn a" in t2*) |
1600 | 15 |
and lts = |
1602
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
16 |
Lnil |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
17 |
| Lcons "name" "trm" "lts" |
1600 | 18 |
binder |
19 |
bn |
|
20 |
where |
|
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
21 |
"bn Lnil = []" |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
22 |
| "bn (Lcons x t l) = (atom x) # (bn l)" |
1600 | 23 |
|
1731
3832a31a73b1
a test with let having multiple bodies
Christian Urban <urbanc@in.tum.de>
parents:
1685
diff
changeset
|
24 |
|
3832a31a73b1
a test with let having multiple bodies
Christian Urban <urbanc@in.tum.de>
parents:
1685
diff
changeset
|
25 |
thm alpha_trm_raw_alpha_lts_raw_alpha_bn_raw.intros |
3832a31a73b1
a test with let having multiple bodies
Christian Urban <urbanc@in.tum.de>
parents:
1685
diff
changeset
|
26 |
|
1600 | 27 |
thm trm_lts.fv |
28 |
thm trm_lts.eq_iff |
|
29 |
thm trm_lts.bn |
|
30 |
thm trm_lts.perm |
|
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
31 |
thm trm_lts.induct[no_vars] |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
32 |
thm trm_lts.inducts[no_vars] |
1600 | 33 |
thm trm_lts.distinct |
1731
3832a31a73b1
a test with let having multiple bodies
Christian Urban <urbanc@in.tum.de>
parents:
1685
diff
changeset
|
34 |
(*thm trm_lts.supp*) |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
35 |
thm trm_lts.fv[simplified trm_lts.supp(1-2)] |
1600 | 36 |
|
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
37 |
|
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
38 |
primrec |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
39 |
permute_bn_raw |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
40 |
where |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
41 |
"permute_bn_raw pi (Lnil_raw) = Lnil_raw" |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
42 |
| "permute_bn_raw pi (Lcons_raw a t l) = Lcons_raw (pi \<bullet> a) t (permute_bn_raw pi l)" |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
43 |
|
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
44 |
quotient_definition |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
45 |
"permute_bn :: perm \<Rightarrow> lts \<Rightarrow> lts" |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
46 |
is |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
47 |
"permute_bn_raw" |
1639
a98d03fb9d53
added experiemental permute_bn
Christian Urban <urbanc@in.tum.de>
parents:
1638
diff
changeset
|
48 |
|
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
49 |
lemma [quot_respect]: "((op =) ===> alpha_lts_raw ===> alpha_lts_raw) permute_bn_raw permute_bn_raw" |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
50 |
apply simp |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
51 |
apply clarify |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
52 |
apply (erule alpha_trm_raw_alpha_lts_raw_alpha_bn_raw.inducts) |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
53 |
apply (rule TrueI)+ |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
54 |
apply simp_all |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
55 |
apply (rule_tac [!] alpha_trm_raw_alpha_lts_raw_alpha_bn_raw.intros) |
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
56 |
apply simp_all |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
57 |
done |
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
58 |
|
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
59 |
lemmas permute_bn = permute_bn_raw.simps[quot_lifted] |
1639
a98d03fb9d53
added experiemental permute_bn
Christian Urban <urbanc@in.tum.de>
parents:
1638
diff
changeset
|
60 |
|
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
61 |
lemma permute_bn_zero: |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
62 |
"permute_bn 0 a = a" |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
63 |
apply(induct a rule: trm_lts.inducts(2)) |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
64 |
apply(rule TrueI)+ |
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
65 |
apply(simp_all add:permute_bn) |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
66 |
done |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
67 |
|
1640
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
68 |
lemma permute_bn_add: |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
69 |
"permute_bn (p + q) a = permute_bn p (permute_bn q a)" |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
70 |
oops |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
71 |
|
1643
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
72 |
lemma permute_bn_alpha_bn: "alpha_bn lts (permute_bn q lts)" |
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
73 |
apply(induct lts rule: trm_lts.inducts(2)) |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
74 |
apply(rule TrueI)+ |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
75 |
apply(simp_all add:permute_bn eqvts trm_lts.eq_iff) |
1643
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
76 |
done |
1641 | 77 |
|
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
78 |
lemma perm_bn: |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
79 |
"p \<bullet> bn l = bn(permute_bn p l)" |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
80 |
apply(induct l rule: trm_lts.inducts(2)) |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
81 |
apply(rule TrueI)+ |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
82 |
apply(simp_all add:permute_bn eqvts) |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
83 |
done |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
84 |
|
1757 | 85 |
lemma fv_perm_bn: |
86 |
"fv_bn l = fv_bn (permute_bn p l)" |
|
87 |
apply(induct l rule: trm_lts.inducts(2)) |
|
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
88 |
apply(rule TrueI)+ |
1757 | 89 |
apply(simp_all add:permute_bn eqvts) |
90 |
done |
|
91 |
||
1643
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
92 |
lemma Lt_subst: |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
93 |
"supp (Abs_lst (bn lts) trm) \<sharp>* q \<Longrightarrow> (Lt lts trm) = Lt (permute_bn q lts) (q \<bullet> trm)" |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
94 |
apply (simp add: trm_lts.eq_iff permute_bn_alpha_bn) |
1643
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
95 |
apply (rule_tac x="q" in exI) |
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
96 |
apply (simp add: alphas) |
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
97 |
apply (simp add: perm_bn[symmetric]) |
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
98 |
apply(rule conjI) |
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
99 |
apply(drule supp_perm_eq) |
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
100 |
apply(simp add: abs_eq_iff) |
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
101 |
apply(simp add: alphas_abs alphas) |
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
102 |
apply(drule conjunct1) |
1643
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
103 |
apply (simp add: trm_lts.supp) |
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
104 |
apply(simp add: supp_abs) |
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2039
diff
changeset
|
105 |
apply (simp add: trm_lts.supp) |
1643
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
106 |
done |
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
107 |
|
953403c5faa0
Showed Let substitution.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1642
diff
changeset
|
108 |
|
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
109 |
lemma fin_bn: |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
110 |
"finite (set (bn l))" |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
111 |
apply(induct l rule: trm_lts.inducts(2)) |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
112 |
apply(simp_all add:permute_bn eqvts) |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
113 |
done |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
114 |
|
1765
9a894c42e80e
more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents:
1759
diff
changeset
|
115 |
thm trm_lts.inducts[no_vars] |
9a894c42e80e
more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents:
1759
diff
changeset
|
116 |
|
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
117 |
lemma |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
118 |
fixes t::trm |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
119 |
and l::lts |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
120 |
and c::"'a::fs" |
1640
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
121 |
assumes a1: "\<And>name c. P1 c (Vr name)" |
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
122 |
and a2: "\<And>trm1 trm2 c. \<lbrakk>\<And>d. P1 d trm1; \<And>d. P1 d trm2\<rbrakk> \<Longrightarrow> P1 c (Ap trm1 trm2)" |
1640
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
123 |
and a3: "\<And>name trm c. \<lbrakk>atom name \<sharp> c; \<And>d. P1 d trm\<rbrakk> \<Longrightarrow> P1 c (Lm name trm)" |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
124 |
and a4: "\<And>lts trm c. \<lbrakk>set (bn lts) \<sharp>* c; \<And>d. P2 d lts; \<And>d. P1 d trm\<rbrakk> \<Longrightarrow> P1 c (Lt lts trm)" |
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
125 |
and a5: "\<And>c. P2 c Lnil" |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
126 |
and a6: "\<And>name trm lts c. \<lbrakk>\<And>d. P1 d trm; \<And>d. P2 d lts\<rbrakk> \<Longrightarrow> P2 c (Lcons name trm lts)" |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
127 |
shows "P1 c t" and "P2 c l" |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
128 |
proof - |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
129 |
have "(\<And>(p::perm) (c::'a::fs). P1 c (p \<bullet> t))" and |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
130 |
b': "(\<And>(p::perm) (q::perm) (c::'a::fs). P2 c (permute_bn p (q \<bullet> l)))" |
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
131 |
apply(induct rule: trm_lts.inducts) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
132 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
133 |
apply(rule a1) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
134 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
135 |
apply(rule a2) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
136 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
137 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
138 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
139 |
apply(subgoal_tac "\<exists>q. (q \<bullet> (atom (p \<bullet> name))) \<sharp> c \<and> supp (Lm (p \<bullet> name) (p \<bullet> trm)) \<sharp>* q") |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
140 |
apply(erule exE) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
141 |
apply(rule_tac t="Lm (p \<bullet> name) (p \<bullet> trm)" |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
142 |
and s="q\<bullet> Lm (p \<bullet> name) (p \<bullet> trm)" in subst) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
143 |
apply(rule supp_perm_eq) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
144 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
145 |
apply(simp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
146 |
apply(rule a3) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
147 |
apply(simp add: atom_eqvt) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
148 |
apply(subst permute_plus[symmetric]) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
149 |
apply(blast) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
150 |
apply(rule at_set_avoiding2_atom) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
151 |
apply(simp add: finite_supp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
152 |
apply(simp add: finite_supp) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
153 |
apply(simp add: fresh_def) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
154 |
apply(simp add: trm_lts.fv[simplified trm_lts.supp]) |
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
155 |
apply(simp) |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
156 |
apply(subgoal_tac "\<exists>q. (q \<bullet> set (bn (p \<bullet> lts))) \<sharp>* c \<and> supp (Abs_lst (bn (p \<bullet> lts)) (p \<bullet> trm)) \<sharp>* q") |
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
157 |
apply(erule exE) |
1641 | 158 |
apply(erule conjE) |
1774
c34347ec7ab3
separated general nominal theory into separate folder
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
159 |
thm Lt_subst |
1641 | 160 |
apply(subst Lt_subst) |
161 |
apply assumption |
|
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
162 |
apply(rule a4) |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
163 |
apply(simp add:perm_bn[symmetric]) |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
164 |
apply(simp add: eqvts) |
1641 | 165 |
apply (simp add: fresh_star_def fresh_def) |
1640
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
166 |
apply(rotate_tac 1) |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
167 |
apply(drule_tac x="q + p" in meta_spec) |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
168 |
apply(simp) |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
169 |
apply(rule at_set_avoiding2) |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
170 |
apply(rule fin_bn) |
1641 | 171 |
apply(simp add: finite_supp) |
172 |
apply(simp add: finite_supp) |
|
1658 | 173 |
apply(simp add: fresh_star_def fresh_def supp_abs) |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
174 |
apply(simp add: eqvts permute_bn) |
1640
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
175 |
apply(rule a5) |
1644
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1643
diff
changeset
|
176 |
apply(simp add: permute_bn) |
1640
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
177 |
apply(rule a6) |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
178 |
apply simp |
cd5a6db05540
trying to prove the string induction for let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1639
diff
changeset
|
179 |
apply simp |
1642
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
180 |
done |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
181 |
then have a: "P1 c (0 \<bullet> t)" by blast |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
182 |
have "P2 c (permute_bn 0 (0 \<bullet> l))" using b' by blast |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
183 |
then show "P1 c t" and "P2 c l" using a permute_bn_zero by simp_all |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
184 |
qed |
06f44d498cef
Only let substitution is left.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1641
diff
changeset
|
185 |
|
1638
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
186 |
|
36798cdbc452
first attempt of strong induction for lets with assignments
Christian Urban <urbanc@in.tum.de>
parents:
1602
diff
changeset
|
187 |
|
1602
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
188 |
lemma lets_bla: |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
189 |
"x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Lt (Lcons x (Vr y) Lnil) (Vr x)) \<noteq> (Lt (Lcons x (Vr z) Lnil) (Vr x))" |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
190 |
by (simp add: trm_lts.eq_iff) |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
191 |
|
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
192 |
lemma lets_ok: |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
193 |
"(Lt (Lcons x (Vr y) Lnil) (Vr x)) = (Lt (Lcons y (Vr y) Lnil) (Vr y))" |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
194 |
apply (simp add: trm_lts.eq_iff) |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
195 |
apply (rule_tac x="(x \<leftrightarrow> y)" in exI) |
2039
39df91a90f87
Move ExLet to NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1774
diff
changeset
|
196 |
apply (simp_all add: alphas eqvts supp_at_base fresh_star_def) |
1602
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
197 |
done |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
198 |
|
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
199 |
lemma lets_ok3: |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
200 |
"x \<noteq> y \<Longrightarrow> |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
201 |
(Lt (Lcons x (Ap (Vr y) (Vr x)) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) \<noteq> |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
202 |
(Lt (Lcons y (Ap (Vr x) (Vr y)) (Lcons x (Vr x) Lnil)) (Ap (Vr x) (Vr y)))" |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
203 |
apply (simp add: alphas trm_lts.eq_iff) |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
204 |
done |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
205 |
|
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
206 |
|
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
207 |
lemma lets_not_ok1: |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
208 |
"x \<noteq> y \<Longrightarrow> |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
209 |
(Lt (Lcons x (Vr x) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) \<noteq> |
1602
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
210 |
(Lt (Lcons y (Vr x) (Lcons x (Vr y) Lnil)) (Ap (Vr x) (Vr y)))" |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1658
diff
changeset
|
211 |
apply (simp add: alphas trm_lts.eq_iff fresh_star_def eqvts) |
1602
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
212 |
done |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
213 |
|
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
214 |
lemma lets_nok: |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
215 |
"x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow> |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
216 |
(Lt (Lcons x (Ap (Vr z) (Vr z)) (Lcons y (Vr z) Lnil)) (Ap (Vr x) (Vr y))) \<noteq> |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
217 |
(Lt (Lcons y (Vr z) (Lcons x (Ap (Vr z) (Vr z)) Lnil)) (Ap (Vr x) (Vr y)))" |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
218 |
apply (simp add: alphas trm_lts.eq_iff fresh_star_def) |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
219 |
done |
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
220 |
|
a7e60da429e2
Move Let properties to ExLet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
221 |
|
1600 | 222 |
end |
223 |
||
224 |
||
225 |