author | Christian Urban <urbanc@in.tum.de> |
Thu, 18 Mar 2010 10:05:36 +0100 | |
changeset 1501 | 7e7dc267ae6b |
parent 1498 | 2ff84b1f551f |
child 1510 | be911e869fde |
permissions | -rw-r--r-- |
1271 | 1 |
theory TySch |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
2 |
imports "Parser" "../Attic/Prove" |
1271 | 3 |
begin |
4 |
||
5 |
atom_decl name |
|
6 |
||
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
7 |
ML {* val _ = cheat_fv_rsp := false *} |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
8 |
ML {* val _ = cheat_const_rsp := false *} |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
9 |
ML {* val _ = cheat_equivp := false *} |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
10 |
ML {* val _ = cheat_fv_eqvt := false *} |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
11 |
ML {* val _ = cheat_alpha_eqvt := false *} |
1271 | 12 |
|
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
13 |
nominal_datatype t = |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
14 |
Var "name" |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
15 |
| Fun "t" "t" |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
16 |
and tyS = |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
17 |
All xs::"name set" ty::"t" bind xs in ty |
1271 | 18 |
|
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
19 |
thm t_tyS_fv |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
20 |
thm t_tyS_eq_iff |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
21 |
thm t_tyS_bn |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
22 |
thm t_tyS_perm |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
23 |
thm t_tyS_induct |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
24 |
thm t_tyS_distinct |
1430
ccbcebef56f3
Trying to prove atom_image_fresh_swap
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
25 |
|
1271 | 26 |
lemma |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
27 |
shows "All {a, b} (Fun (Var a) (Var b)) = All {b, a} (Fun (Var a) (Var b))" |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
28 |
apply(simp add: t_tyS_eq_iff) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
29 |
apply(rule_tac x="0::perm" in exI) |
1271 | 30 |
apply(simp add: alpha_gen) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
31 |
apply(auto) |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
32 |
apply(simp add: fresh_star_def fresh_zero_perm) |
1271 | 33 |
done |
34 |
||
35 |
lemma |
|
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
36 |
shows "All {a, b} (Fun (Var a) (Var b)) = All {a, b} (Fun (Var b) (Var a))" |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
37 |
apply(simp add: t_tyS_eq_iff) |
1271 | 38 |
apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
39 |
apply(simp add: alpha_gen t_tyS_fv fresh_star_def t_tyS_perm eqvts) |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
40 |
apply auto |
1271 | 41 |
done |
42 |
||
43 |
lemma |
|
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
44 |
shows "All {a, b, c} (Fun (Var a) (Var b)) = All {a, b} (Fun (Var a) (Var b))" |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
45 |
apply(simp add: t_tyS_eq_iff) |
1271 | 46 |
apply(rule_tac x="0::perm" in exI) |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
47 |
apply(simp add: alpha_gen t_tyS_fv fresh_star_def t_tyS_perm eqvts t_tyS_eq_iff) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
48 |
oops |
1271 | 49 |
|
50 |
lemma |
|
51 |
assumes a: "a \<noteq> b" |
|
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
52 |
shows "\<not>(All {a, b} (Fun (Var a) (Var b)) = All {c} (Fun (Var c) (Var c)))" |
1271 | 53 |
using a |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
54 |
apply(simp add: t_tyS_eq_iff) |
1271 | 55 |
apply(clarify) |
1496
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1477
diff
changeset
|
56 |
apply(simp add: alpha_gen t_tyS_fv fresh_star_def t_tyS_perm eqvts t_tyS_eq_iff) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
57 |
apply auto |
1271 | 58 |
done |
59 |
||
60 |
end |