Nominal/LFex.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 05 Mar 2010 17:09:48 +0100
changeset 1348 2e2a3cd58f64
parent 1344 b320da14e63c
child 1360 c54cb3f7ac70
permissions -rw-r--r--
Ported LF to the parser interface.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
985
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory LFex
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
     2
imports "Parser"
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
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
     8
ML {* restricted_nominal := false *}
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
     9
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    10
nominal_datatype kind =
985
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
    Type
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    12
  | KPi "ty" n::"name" k::"kind" bind n in k
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    13
and ty =
985
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
    TConst "ident"
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    15
  | TApp "ty" "trm"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    16
  | TPi "ty" n::"name" t::"ty" bind n in t
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    17
and trm =
985
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
    Const "ident"
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
  | Var "name"
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    20
  | App "trm" "trm"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    21
  | Lam "ty" n::"name" t::"trm" bind n in t
994
333c24bd595d More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 993
diff changeset
    22
1244
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    23
lemma supports:
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    24
  "{} supports Type"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    25
  "(supp (atom i)) supports (TConst i)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    26
  "(supp A \<union> supp M) supports (TApp A M)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    27
  "(supp (atom i)) supports (Const i)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    28
  "(supp (atom x)) supports (Var x)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    29
  "(supp M \<union> supp N) supports (App M N)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    30
  "(supp ty \<union> supp (atom na) \<union> supp ki) supports (KPi ty na ki)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    31
  "(supp ty \<union> supp (atom na) \<union> supp ty2) supports (TPi ty na ty2)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    32
  "(supp ty \<union> supp (atom na) \<union> supp trm) supports (Lam ty na trm)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    33
apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh kind_ty_trm_perm)
1244
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    34
apply(rule_tac [!] allI)+
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    35
apply(rule_tac [!] impI)
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    36
apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    37
apply(simp_all add: fresh_atom)
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    38
done
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    39
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    40
lemma kind_ty_trm_fs:
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    41
  "finite (supp (x\<Colon>kind)) \<and> finite (supp (y\<Colon>ty)) \<and> finite (supp (z\<Colon>trm))"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    42
apply(induct rule: kind_ty_trm_induct)
1244
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    43
apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    44
apply(simp_all add: supp_atom)
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    45
done
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    46
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    47
instance kind and ty and trm :: fs
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    48
apply(default)
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    49
apply(simp_all only: kind_ty_trm_fs)
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    50
done
605a0ebe87da Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1243
diff changeset
    51
1344
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    52
lemma ex_out: 
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    53
  "(\<exists>x. Z x \<and> Q) = (Q \<and> (\<exists>x. Z x))"
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    54
  "(\<exists>x. Q \<and> Z x) = (Q \<and> (\<exists>x. Z x))"
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    55
  "(\<exists>x. P x \<and> Q \<and> Z x) = (Q \<and> (\<exists>x. P x \<and> Z x))"
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    56
  "(\<exists>x. Q \<and> P x \<and> Z x) = (Q \<and> (\<exists>x. P x \<and> Z x))"
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    57
apply (blast)+
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    58
done
b320da14e63c Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1310
diff changeset
    59
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    60
lemma Collect_neg_conj: "{x. \<not>(P x \<and> Q x)} = {x. \<not>(P x)} \<union> {x. \<not>(Q x)}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    61
by (simp add: Collect_imp_eq Collect_neg_eq[symmetric])
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    62
1245
18310dff4e3a Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1244
diff changeset
    63
lemma supp_eqs:
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    64
  "supp Type = {}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    65
  "supp rkind = fv_kind rkind \<Longrightarrow> supp (KPi rty name rkind) = supp rty \<union> supp (Abs {atom name} rkind)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    66
  "supp (TConst i) = {atom i}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    67
  "supp (TApp A M) = supp A \<union> supp M"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    68
  "supp rty2 = fv_ty rty2 \<Longrightarrow> supp (TPi rty1 name rty2) = supp rty1 \<union> supp (Abs {atom name} rty2)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    69
  "supp (Const i) = {atom i}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    70
  "supp (Var x) = {atom x}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    71
  "supp (App M N) = supp M \<union> supp N"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    72
  "supp rtrm = fv_trm rtrm \<Longrightarrow> supp (Lam rty name rtrm) = supp rty \<union> supp (Abs {atom name} rtrm)"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    73
  apply(simp_all (no_asm) add: supp_def permute_set_eq atom_eqvt kind_ty_trm_perm)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    74
  apply(simp_all only: kind_ty_trm_inject Abs_eq_iff alpha_gen)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    75
  apply(simp_all only: ex_out)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    76
  apply(simp_all only: eqvts[symmetric])
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    77
  apply(simp_all only: Collect_neg_conj)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    78
  apply(simp_all only: supp_at_base[simplified supp_def] Un_commute Un_assoc)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    79
  apply(simp_all add: Collect_imp_eq Collect_neg_eq[symmetric] Un_commute Un_assoc)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    80
  apply(simp_all add: Un_left_commute)
1245
18310dff4e3a Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1244
diff changeset
    81
  done
1002
3f227ed7e3e5 More proofs in the LF example.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 997
diff changeset
    82
994
333c24bd595d More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 993
diff changeset
    83
lemma supp_fv:
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    84
  "supp t1 = fv_kind t1 \<and> supp t2 = fv_ty t2 \<and> supp t3 = fv_trm t3"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    85
  apply(induct rule: kind_ty_trm_induct)
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    86
  apply(simp_all (no_asm) only: supp_eqs kind_ty_trm_fv)
1245
18310dff4e3a Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1244
diff changeset
    87
  apply(simp_all)
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    88
  apply(simp_all add: supp_eqs)
1245
18310dff4e3a Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1244
diff changeset
    89
  apply(simp_all add: supp_Abs)
18310dff4e3a Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1244
diff changeset
    90
  done
994
333c24bd595d More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 993
diff changeset
    91
1234
1240d5eb275d LF renaming part1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1219
diff changeset
    92
lemma supp_rkind_rty_rtrm:
1348
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    93
  "supp Type = {}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    94
  "supp (KPi A x K) = supp A \<union> (supp K - {atom x})"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    95
  "supp (TConst i) = {atom i}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    96
  "supp (TApp A M) = supp A \<union> supp M"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    97
  "supp (TPi A x B) = supp A \<union> (supp B - {atom x})"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    98
  "supp (Const i) = {atom i}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
    99
  "supp (Var x) = {atom x}"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
   100
  "supp (App M N) = supp M \<union> supp N"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
   101
  "supp (Lam A x M) = supp A \<union> (supp M - {atom x})"
2e2a3cd58f64 Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1344
diff changeset
   102
apply (simp_all add: supp_fv kind_ty_trm_fv)
994
333c24bd595d More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 993
diff changeset
   103
985
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
end
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
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
ef8a2b0b237a Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108