Nominal/Ex/CoreHaskell.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 28 Dec 2010 00:20:50 +0000
changeset 2629 ffb5a181844b
parent 2598 b136721eedb2
child 2630 8268b277d240
permissions -rw-r--r--
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2083
9568f9f31822 tuned file names for examples
Christian Urban <urbanc@in.tum.de>
parents: 2064
diff changeset
     1
theory CoreHaskell
2454
9ffee4eb1ae1 renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 2436
diff changeset
     2
imports "../Nominal2"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
2308
387fcbd33820 fixed problem with bn_info
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
     5
(* Core Haskell *)
2100
705dc7532ee3 added comment about bind_set
Christian Urban <urbanc@in.tum.de>
parents: 2095
diff changeset
     6
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
atom_decl var
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
     8
atom_decl cvar
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
atom_decl tvar
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    11
nominal_datatype core_haskell: 
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    12
 tkind =
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
  KStar
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
| KFun "tkind" "tkind"
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
and ckind =
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  CKEq "ty" "ty"
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
and ty =
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
  TVar "tvar"
1684
b9e4c812d2c6 Core Haskell can now use proper strings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1667
diff changeset
    19
| TC "string"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
| TApp "ty" "ty"
1684
b9e4c812d2c6 Core Haskell can now use proper strings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1667
diff changeset
    21
| TFun "string" "ty_lst"
2424
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    22
| TAll tv::"tvar" "tkind" T::"ty"  bind (set) tv in T
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    23
| TEq "ckind" "ty"
1606
75403378068b Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1601
diff changeset
    24
and ty_lst =
75403378068b Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1601
diff changeset
    25
  TsNil
75403378068b Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1601
diff changeset
    26
| TsCons "ty" "ty_lst"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
and co =
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1698
diff changeset
    28
  CVar "cvar"
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    29
| CConst "string"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
| CApp "co" "co"
1684
b9e4c812d2c6 Core Haskell can now use proper strings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1667
diff changeset
    31
| CFun "string" "co_lst"
2424
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    32
| CAll cv::"cvar" "ckind" C::"co"  bind (set) cv in C
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    33
| CEq "ckind" "co"
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    34
| CRefl "ty"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
| CSym "co"
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    36
| CCir "co" "co"
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    37
| CAt "co" "ty"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
| CLeft "co"
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
| CRight "co"
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    40
| CSim "co" "co"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
| CRightc "co"
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
| CLeftc "co"
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
| CCoe "co" "co"
1606
75403378068b Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1601
diff changeset
    44
and co_lst =
75403378068b Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1601
diff changeset
    45
  CsNil
75403378068b Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1601
diff changeset
    46
| CsCons "co" "co_lst"
1609
c9bc3b61046c Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1606
diff changeset
    47
and trm =
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
  Var "var"
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1698
diff changeset
    49
| K "string"
2424
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    50
| LAMT tv::"tvar" "tkind" t::"trm" bind (set) tv in t
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    51
| LAMC cv::"cvar" "ckind" t::"trm" bind (set) cv in t
1698
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    52
| AppT "trm" "ty"
69472e74bd3b Update according to paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1695
diff changeset
    53
| AppC "trm" "co"
2424
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    54
| Lam v::"var" "ty" t::"trm"       bind (set) v in t
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
| App "trm" "trm"
2424
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    56
| Let x::"var" "ty" "trm" t::"trm" bind (set) x in t
1609
c9bc3b61046c Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1606
diff changeset
    57
| Case "trm" "assoc_lst"
c9bc3b61046c Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1606
diff changeset
    58
| Cast "trm" "ty"                  --"ty is supposed to be a coercion type only"
c9bc3b61046c Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1606
diff changeset
    59
and assoc_lst =
c9bc3b61046c Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1606
diff changeset
    60
  ANil
c9bc3b61046c Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1606
diff changeset
    61
| ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
and pat =
1700
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    63
  Kpat "string" "tvars" "cvars" "vars"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    64
and vars =
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    65
  VsNil
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    66
| VsCons "var" "ty" "vars"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    67
and tvars =
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    68
  TvsNil
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    69
| TvsCons "tvar" "tkind" "tvars"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    70
and cvars =
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    71
  CvsNil
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    72
| CvsCons "cvar" "ckind" "cvars"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
binder
1695
e42ee5947b5c Changed to Lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1688
diff changeset
    74
    bv :: "pat \<Rightarrow> atom list"
1700
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    75
and bv_vs :: "vars \<Rightarrow> atom list"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    76
and bv_tvs :: "tvars \<Rightarrow> atom list"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    77
and bv_cvs :: "cvars \<Rightarrow> atom list"
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
where
1700
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    79
  "bv (Kpat s tvts tvcs vs) = append (bv_tvs tvts) (append (bv_cvs tvcs) (bv_vs vs))"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    80
| "bv_vs VsNil = []"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    81
| "bv_vs (VsCons v k t) = (atom v) # bv_vs t"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    82
| "bv_tvs TvsNil = []"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    83
| "bv_tvs (TvsCons v k t) = (atom v) # bv_tvs t"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
    84
| "bv_cvs CvsNil = []"
2213
231a20534950 improved abstract, some tuning
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
    85
| "bv_cvs (CvsCons v k t) = (atom v) # bv_cvs t"
231a20534950 improved abstract, some tuning
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
    86
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    87
(* generated theorems *)
2406
428d9cb9a243 can also lift the various eqvt lemmas for bn, fv, fv_bn and size
Christian Urban <urbanc@in.tum.de>
parents: 2405
diff changeset
    88
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2593
diff changeset
    89
thm core_haskell.permute_bn
2593
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    90
thm core_haskell.perm_bn_alpha
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    91
thm core_haskell.perm_bn_simps
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    92
thm core_haskell.bn_finite
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    93
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    94
thm core_haskell.distinct
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    95
thm core_haskell.induct
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    96
thm core_haskell.exhaust
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    97
thm core_haskell.fv_defs
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    98
thm core_haskell.bn_defs
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    99
thm core_haskell.perm_simps
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
   100
thm core_haskell.eq_iff
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
   101
thm core_haskell.fv_bn_eqvt
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
   102
thm core_haskell.size_eqvt
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
   103
thm core_haskell.supp
2339
1e0b3992189c more quotient-definitions
Christian Urban <urbanc@in.tum.de>
parents: 2336
diff changeset
   104
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   105
lemma strong_induction_principle:
2629
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   106
  fixes c::"'a::fs"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   107
  assumes a01: "\<And>b. P1 b KStar"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   108
  and a02: "\<And>tkind1 tkind2 b. \<lbrakk>\<And>c. P1 c tkind1; \<And>c. P1 c tkind2\<rbrakk> \<Longrightarrow> P1 b (KFun tkind1 tkind2)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   109
  and a03: "\<And>ty1 ty2 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2\<rbrakk> \<Longrightarrow> P2 b (CKEq ty1 ty2)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   110
  and a04: "\<And>tvar b. P3 b (TVar tvar)"
1684
b9e4c812d2c6 Core Haskell can now use proper strings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1667
diff changeset
   111
  and a05: "\<And>string b. P3 b (TC string)"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   112
  and a06: "\<And>ty1 ty2 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2\<rbrakk> \<Longrightarrow> P3 b (TApp ty1 ty2)"
1684
b9e4c812d2c6 Core Haskell can now use proper strings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1667
diff changeset
   113
  and a07: "\<And>string ty_lst b. \<lbrakk>\<And>c. P4 c ty_lst\<rbrakk> \<Longrightarrow> P3 b (TFun string ty_lst)"
1632
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   114
  and a08: "\<And>tvar tkind ty b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P3 c ty; atom tvar \<sharp> b\<rbrakk>
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   115
    \<Longrightarrow> P3 b (TAll tvar tkind ty)"
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   116
  and a09: "\<And>ck ty b. \<lbrakk>\<And>c. P2 c ck; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P3 b (TEq ck ty)"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   117
  and a10: "\<And>b. P4 b TsNil"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   118
  and a11: "\<And>ty ty_lst b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P4 c ty_lst\<rbrakk> \<Longrightarrow> P4 b (TsCons ty ty_lst)"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   119
  and a12: "\<And>string b. P5 b (CVar string)"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   120
  and a12a:"\<And>str b. P5 b (CConst str)"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   121
  and a13: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CApp co1 co2)"
1684
b9e4c812d2c6 Core Haskell can now use proper strings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1667
diff changeset
   122
  and a14: "\<And>string co_lst b. \<lbrakk>\<And>c. P6 c co_lst\<rbrakk> \<Longrightarrow> P5 b (CFun string co_lst)"
1632
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   123
  and a15: "\<And>tvar ckind co b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P5 c co; atom tvar \<sharp> b\<rbrakk>
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   124
    \<Longrightarrow> P5 b (CAll tvar ckind co)"
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   125
  and a16: "\<And>ck co b. \<lbrakk>\<And>c. P2 c ck; \<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CEq ck co)"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   126
  and a17: "\<And>ty b. \<lbrakk>\<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P5 b (CRefl ty)"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   127
  and a17a: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CSym co)"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   128
  and a18: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CCir co1 co2)"
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   129
  and a18a:"\<And>co ty b. \<lbrakk>\<And>c. P5 c co; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P5 b (CAt co ty)"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   130
  and a19: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CLeft co)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   131
  and a20: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CRight co)"
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   132
  and a21: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CSim co1 co2)"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   133
  and a22: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CRightc co)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   134
  and a23: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CLeftc co)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   135
  and a24: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CCoe co1 co2)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   136
  and a25: "\<And>b. P6 b CsNil"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   137
  and a26: "\<And>co co_lst b. \<lbrakk>\<And>c. P5 c co; \<And>c. P6 c co_lst\<rbrakk> \<Longrightarrow> P6 b (CsCons co co_lst)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   138
  and a27: "\<And>var b. P7 b (Var var)"
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   139
  and a28: "\<And>string b. P7 b (K string)"
1632
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   140
  and a29: "\<And>tvar tkind trm b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P7 c trm; atom tvar \<sharp> b\<rbrakk>
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   141
    \<Longrightarrow> P7 b (LAMT tvar tkind trm)"
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   142
  and a30: "\<And>tvar ckind trm b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P7 c trm; atom tvar \<sharp> b\<rbrakk>
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   143
    \<Longrightarrow> P7 b (LAMC tvar ckind trm)"
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   144
  and a31: "\<And>trm ty b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P7 b (AppT trm ty)"
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   145
  and a31a:"\<And>trm co b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P5 c co\<rbrakk> \<Longrightarrow> P7 b (AppC trm co)"
1632
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   146
  and a32: "\<And>var ty trm b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P7 c trm; atom var \<sharp> b\<rbrakk> \<Longrightarrow> P7 b (Lam var ty trm)"
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   147
  and a33: "\<And>trm1 trm2 b. \<lbrakk>\<And>c. P7 c trm1; \<And>c. P7 c trm2\<rbrakk> \<Longrightarrow> P7 b (App trm1 trm2)"
1632
68c8666453f7 Started proving strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1631
diff changeset
   148
  and a34: "\<And>var ty trm1 trm2 b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P7 c trm1; \<And>c. P7 c trm2; atom var \<sharp> b\<rbrakk>
1631
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   149
    \<Longrightarrow> P7 b (Let var ty trm1 trm2)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   150
  and a35: "\<And>trm assoc_lst b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P8 c assoc_lst\<rbrakk> \<Longrightarrow> P7 b (Case trm assoc_lst)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   151
  and a36: "\<And>trm ty b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P7 b (Cast trm ty)"
e94bfef17bb8 stating the strong induction; further.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1630
diff changeset
   152
  and a37: "\<And>b. P8 b ANil"
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
   153
  and a38: "\<And>pt trm assoc_lst b. \<lbrakk>\<And>c. P9 c pt; \<And>c. P7 c trm; \<And>c. P8 c assoc_lst; set (bv (pt)) \<sharp>* b\<rbrakk>
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
   154
    \<Longrightarrow> P8 b (ACons pt trm assoc_lst)"
1700
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   155
  and a39: "\<And>string tvars cvars vars b. \<lbrakk>\<And>c. P11 c tvars; \<And>c. P12 c cvars; \<And>c. P10 c vars\<rbrakk>
1701
58abc09d6f9c Updated strong induction to modified definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1700
diff changeset
   156
    \<Longrightarrow> P9 b (Kpat string tvars cvars vars)"
1700
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   157
  and a40: "\<And>b. P10 b VsNil"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   158
  and a41: "\<And>var ty vars b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P10 c vars\<rbrakk> \<Longrightarrow> P10 b (VsCons var ty vars)"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   159
  and a42: "\<And>b. P11 b TvsNil"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   160
  and a43: "\<And>tvar tkind tvars b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P11 c tvars\<rbrakk>
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   161
    \<Longrightarrow> P11 b (TvsCons tvar tkind tvars)"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   162
  and a44: "\<And>b. P12 b CvsNil"
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   163
  and a45: "\<And>tvar ckind cvars b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P12 c cvars\<rbrakk>
b5141d1ab24f Initial renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1699
diff changeset
   164
    \<Longrightarrow> P12 b (CvsCons tvar ckind cvars)"
2629
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   165
  shows "P1 c tkind"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   166
        "P2 c ckind"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   167
        "P3 c ty"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   168
        "P4 c ty_lst"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   169
        "P5 c co"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   170
        "P6 c co_lst"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   171
        "P7 c trm"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   172
        "P8 c assoc_lst"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   173
        "P9 c pt"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   174
        "P10 c vars"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   175
        "P11 c tvars"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   176
        "P12 c cvars"
ffb5a181844b proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
   177
oops
1868
26c0c35641cb equivariance for alpha_raw in CoreHaskell is automatically derived
Christian Urban <urbanc@in.tum.de>
parents: 1867
diff changeset
   178
1601
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
end
5f0bb35114c3 Added missing file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180