src/Cube/Example.thy
changeset 45242 401f91ed8a93
parent 42814 5af15f1e2ef6
child 58617 4f169d2cf6f3
equal deleted inserted replaced
45241:87950f752099 45242:401f91ed8a93
    28 *}
    28 *}
    29 
    29 
    30 
    30 
    31 subsection {* Simple types *}
    31 subsection {* Simple types *}
    32 
    32 
    33 schematic_lemma "A:* |- A->A : ?T"
    33 schematic_lemma "A:* \<turnstile> A\<rightarrow>A : ?T"
    34   by (depth_solve rules)
    34   by (depth_solve rules)
    35 
    35 
    36 schematic_lemma "A:* |- Lam a:A. a : ?T"
    36 schematic_lemma "A:* \<turnstile> \<Lambda> a:A. a : ?T"
    37   by (depth_solve rules)
    37   by (depth_solve rules)
    38 
    38 
    39 schematic_lemma "A:* B:* b:B |- Lam x:A. b : ?T"
    39 schematic_lemma "A:* B:* b:B \<turnstile> \<Lambda> x:A. b : ?T"
    40   by (depth_solve rules)
    40   by (depth_solve rules)
    41 
    41 
    42 schematic_lemma "A:* b:A |- (Lam a:A. a)^b: ?T"
    42 schematic_lemma "A:* b:A \<turnstile> (\<Lambda> a:A. a)^b: ?T"
    43   by (depth_solve rules)
    43   by (depth_solve rules)
    44 
    44 
    45 schematic_lemma "A:* B:* c:A b:B |- (Lam x:A. b)^ c: ?T"
    45 schematic_lemma "A:* B:* c:A b:B \<turnstile> (\<Lambda> x:A. b)^ c: ?T"
    46   by (depth_solve rules)
    46   by (depth_solve rules)
    47 
    47 
    48 schematic_lemma "A:* B:* |- Lam a:A. Lam b:B. a : ?T"
    48 schematic_lemma "A:* B:* \<turnstile> \<Lambda> a:A. \<Lambda> b:B. a : ?T"
    49   by (depth_solve rules)
    49   by (depth_solve rules)
    50 
    50 
    51 
    51 
    52 subsection {* Second-order types *}
    52 subsection {* Second-order types *}
    53 
    53 
    54 schematic_lemma (in L2) "|- Lam A:*. Lam a:A. a : ?T"
    54 schematic_lemma (in L2) "\<turnstile> \<Lambda> A:*. \<Lambda> a:A. a : ?T"
    55   by (depth_solve rules)
    55   by (depth_solve rules)
    56 
    56 
    57 schematic_lemma (in L2) "A:* |- (Lam B:*.Lam b:B. b)^A : ?T"
    57 schematic_lemma (in L2) "A:* \<turnstile> (\<Lambda> B:*.\<Lambda> b:B. b)^A : ?T"
    58   by (depth_solve rules)
    58   by (depth_solve rules)
    59 
    59 
    60 schematic_lemma (in L2) "A:* b:A |- (Lam B:*.Lam b:B. b) ^ A ^ b: ?T"
    60 schematic_lemma (in L2) "A:* b:A \<turnstile> (\<Lambda> B:*.\<Lambda> b:B. b) ^ A ^ b: ?T"
    61   by (depth_solve rules)
    61   by (depth_solve rules)
    62 
    62 
    63 schematic_lemma (in L2) "|- Lam B:*.Lam a:(Pi A:*.A).a ^ ((Pi A:*.A)->B) ^ a: ?T"
    63 schematic_lemma (in L2) "\<turnstile> \<Lambda> B:*.\<Lambda> a:(\<Pi> A:*.A).a ^ ((\<Pi> A:*.A)\<rightarrow>B) ^ a: ?T"
    64   by (depth_solve rules)
    64   by (depth_solve rules)
    65 
    65 
    66 
    66 
    67 subsection {* Weakly higher-order propositional logic *}
    67 subsection {* Weakly higher-order propositional logic *}
    68 
    68 
    69 schematic_lemma (in Lomega) "|- Lam A:*.A->A : ?T"
    69 schematic_lemma (in Lomega) "\<turnstile> \<Lambda> A:*.A\<rightarrow>A : ?T"
    70   by (depth_solve rules)
    70   by (depth_solve rules)
    71 
    71 
    72 schematic_lemma (in Lomega) "B:* |- (Lam A:*.A->A) ^ B : ?T"
    72 schematic_lemma (in Lomega) "B:* \<turnstile> (\<Lambda> A:*.A\<rightarrow>A) ^ B : ?T"
    73   by (depth_solve rules)
    73   by (depth_solve rules)
    74 
    74 
    75 schematic_lemma (in Lomega) "B:* b:B |- (Lam y:B. b): ?T"
    75 schematic_lemma (in Lomega) "B:* b:B \<turnstile> (\<Lambda> y:B. b): ?T"
    76   by (depth_solve rules)
    76   by (depth_solve rules)
    77 
    77 
    78 schematic_lemma (in Lomega) "A:* F:*->* |- F^(F^A): ?T"
    78 schematic_lemma (in Lomega) "A:* F:*\<rightarrow>* \<turnstile> F^(F^A): ?T"
    79   by (depth_solve rules)
    79   by (depth_solve rules)
    80 
    80 
    81 schematic_lemma (in Lomega) "A:* |- Lam F:*->*.F^(F^A): ?T"
    81 schematic_lemma (in Lomega) "A:* \<turnstile> \<Lambda> F:*\<rightarrow>*.F^(F^A): ?T"
    82   by (depth_solve rules)
    82   by (depth_solve rules)
    83 
    83 
    84 
    84 
    85 subsection {* LP *}
    85 subsection {* LP *}
    86 
    86 
    87 schematic_lemma (in LP) "A:* |- A -> * : ?T"
    87 schematic_lemma (in LP) "A:* \<turnstile> A \<rightarrow> * : ?T"
    88   by (depth_solve rules)
    88   by (depth_solve rules)
    89 
    89 
    90 schematic_lemma (in LP) "A:* P:A->* a:A |- P^a: ?T"
    90 schematic_lemma (in LP) "A:* P:A\<rightarrow>* a:A \<turnstile> P^a: ?T"
    91   by (depth_solve rules)
    91   by (depth_solve rules)
    92 
    92 
    93 schematic_lemma (in LP) "A:* P:A->A->* a:A |- Pi a:A. P^a^a: ?T"
    93 schematic_lemma (in LP) "A:* P:A\<rightarrow>A\<rightarrow>* a:A \<turnstile> \<Pi> a:A. P^a^a: ?T"
    94   by (depth_solve rules)
    94   by (depth_solve rules)
    95 
    95 
    96 schematic_lemma (in LP) "A:* P:A->* Q:A->* |- Pi a:A. P^a -> Q^a: ?T"
    96 schematic_lemma (in LP) "A:* P:A\<rightarrow>* Q:A\<rightarrow>* \<turnstile> \<Pi> a:A. P^a \<rightarrow> Q^a: ?T"
    97   by (depth_solve rules)
    97   by (depth_solve rules)
    98 
    98 
    99 schematic_lemma (in LP) "A:* P:A->* |- Pi a:A. P^a -> P^a: ?T"
    99 schematic_lemma (in LP) "A:* P:A\<rightarrow>* \<turnstile> \<Pi> a:A. P^a \<rightarrow> P^a: ?T"
   100   by (depth_solve rules)
   100   by (depth_solve rules)
   101 
   101 
   102 schematic_lemma (in LP) "A:* P:A->* |- Lam a:A. Lam x:P^a. x: ?T"
   102 schematic_lemma (in LP) "A:* P:A\<rightarrow>* \<turnstile> \<Lambda> a:A. \<Lambda> x:P^a. x: ?T"
   103   by (depth_solve rules)
   103   by (depth_solve rules)
   104 
   104 
   105 schematic_lemma (in LP) "A:* P:A->* Q:* |- (Pi a:A. P^a->Q) -> (Pi a:A. P^a) -> Q : ?T"
   105 schematic_lemma (in LP) "A:* P:A\<rightarrow>* Q:* \<turnstile> (\<Pi> a:A. P^a\<rightarrow>Q) \<rightarrow> (\<Pi> a:A. P^a) \<rightarrow> Q : ?T"
   106   by (depth_solve rules)
   106   by (depth_solve rules)
   107 
   107 
   108 schematic_lemma (in LP) "A:* P:A->* Q:* a0:A |-
   108 schematic_lemma (in LP) "A:* P:A\<rightarrow>* Q:* a0:A \<turnstile>
   109         Lam x:Pi a:A. P^a->Q. Lam y:Pi a:A. P^a. x^a0^(y^a0): ?T"
   109         \<Lambda> x:\<Pi> a:A. P^a\<rightarrow>Q. \<Lambda> y:\<Pi> a:A. P^a. x^a0^(y^a0): ?T"
   110   by (depth_solve rules)
   110   by (depth_solve rules)
   111 
   111 
   112 
   112 
   113 subsection {* Omega-order types *}
   113 subsection {* Omega-order types *}
   114 
   114 
   115 schematic_lemma (in L2) "A:* B:* |- Pi C:*.(A->B->C)->C : ?T"
   115 schematic_lemma (in L2) "A:* B:* \<turnstile> \<Pi> C:*.(A\<rightarrow>B\<rightarrow>C)\<rightarrow>C : ?T"
   116   by (depth_solve rules)
   116   by (depth_solve rules)
   117 
   117 
   118 schematic_lemma (in Lomega2) "|- Lam A:*.Lam B:*.Pi C:*.(A->B->C)->C : ?T"
   118 schematic_lemma (in Lomega2) "\<turnstile> \<Lambda> A:*.\<Lambda> B:*.\<Pi> C:*.(A\<rightarrow>B\<rightarrow>C)\<rightarrow>C : ?T"
   119   by (depth_solve rules)
   119   by (depth_solve rules)
   120 
   120 
   121 schematic_lemma (in Lomega2) "|- Lam A:*.Lam B:*.Lam x:A. Lam y:B. x : ?T"
   121 schematic_lemma (in Lomega2) "\<turnstile> \<Lambda> A:*.\<Lambda> B:*.\<Lambda> x:A. \<Lambda> y:B. x : ?T"
   122   by (depth_solve rules)
   122   by (depth_solve rules)
   123 
   123 
   124 schematic_lemma (in Lomega2) "A:* B:* |- ?p : (A->B) -> ((B->Pi P:*.P)->(A->Pi P:*.P))"
   124 schematic_lemma (in Lomega2) "A:* B:* \<turnstile> ?p : (A\<rightarrow>B) \<rightarrow> ((B\<rightarrow>\<Pi> P:*.P)\<rightarrow>(A\<rightarrow>\<Pi> P:*.P))"
   125   apply (strip_asms rules)
   125   apply (strip_asms rules)
   126   apply (rule lam_ss)
   126   apply (rule lam_ss)
   127     apply (depth_solve1 rules)
   127     apply (depth_solve1 rules)
   128    prefer 2
   128    prefer 2
   129    apply (depth_solve1 rules)
   129    apply (depth_solve1 rules)
   143   done
   143   done
   144 
   144 
   145 
   145 
   146 subsection {* Second-order Predicate Logic *}
   146 subsection {* Second-order Predicate Logic *}
   147 
   147 
   148 schematic_lemma (in LP2) "A:* P:A->* |- Lam a:A. P^a->(Pi A:*.A) : ?T"
   148 schematic_lemma (in LP2) "A:* P:A\<rightarrow>* \<turnstile> \<Lambda> a:A. P^a\<rightarrow>(\<Pi> A:*.A) : ?T"
   149   by (depth_solve rules)
   149   by (depth_solve rules)
   150 
   150 
   151 schematic_lemma (in LP2) "A:* P:A->A->* |-
   151 schematic_lemma (in LP2) "A:* P:A\<rightarrow>A\<rightarrow>* \<turnstile>
   152     (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P : ?T"
   152     (\<Pi> a:A. \<Pi> b:A. P^a^b\<rightarrow>P^b^a\<rightarrow>\<Pi> P:*.P) \<rightarrow> \<Pi> a:A. P^a^a\<rightarrow>\<Pi> P:*.P : ?T"
   153   by (depth_solve rules)
   153   by (depth_solve rules)
   154 
   154 
   155 schematic_lemma (in LP2) "A:* P:A->A->* |-
   155 schematic_lemma (in LP2) "A:* P:A\<rightarrow>A\<rightarrow>* \<turnstile>
   156     ?p: (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P"
   156     ?p: (\<Pi> a:A. \<Pi> b:A. P^a^b\<rightarrow>P^b^a\<rightarrow>\<Pi> P:*.P) \<rightarrow> \<Pi> a:A. P^a^a\<rightarrow>\<Pi> P:*.P"
   157   -- {* Antisymmetry implies irreflexivity: *}
   157   -- {* Antisymmetry implies irreflexivity: *}
   158   apply (strip_asms rules)
   158   apply (strip_asms rules)
   159   apply (rule lam_ss)
   159   apply (rule lam_ss)
   160     apply (depth_solve1 rules)
   160     apply (depth_solve1 rules)
   161    prefer 2
   161    prefer 2
   172   done
   172   done
   173 
   173 
   174 
   174 
   175 subsection {* LPomega *}
   175 subsection {* LPomega *}
   176 
   176 
   177 schematic_lemma (in LPomega) "A:* |- Lam P:A->A->*.Lam a:A. P^a^a : ?T"
   177 schematic_lemma (in LPomega) "A:* \<turnstile> \<Lambda> P:A\<rightarrow>A\<rightarrow>*.\<Lambda> a:A. P^a^a : ?T"
   178   by (depth_solve rules)
   178   by (depth_solve rules)
   179 
   179 
   180 schematic_lemma (in LPomega) "|- Lam A:*.Lam P:A->A->*.Lam a:A. P^a^a : ?T"
   180 schematic_lemma (in LPomega) "\<turnstile> \<Lambda> A:*.\<Lambda> P:A\<rightarrow>A\<rightarrow>*.\<Lambda> a:A. P^a^a : ?T"
   181   by (depth_solve rules)
   181   by (depth_solve rules)
   182 
   182 
   183 
   183 
   184 subsection {* Constructions *}
   184 subsection {* Constructions *}
   185 
   185 
   186 schematic_lemma (in CC) "|- Lam A:*.Lam P:A->*.Lam a:A. P^a->Pi P:*.P: ?T"
   186 schematic_lemma (in CC) "\<turnstile> \<Lambda> A:*.\<Lambda> P:A\<rightarrow>*.\<Lambda> a:A. P^a\<rightarrow>\<Pi> P:*.P: ?T"
   187   by (depth_solve rules)
   187   by (depth_solve rules)
   188 
   188 
   189 schematic_lemma (in CC) "|- Lam A:*.Lam P:A->*.Pi a:A. P^a: ?T"
   189 schematic_lemma (in CC) "\<turnstile> \<Lambda> A:*.\<Lambda> P:A\<rightarrow>*.\<Pi> a:A. P^a: ?T"
   190   by (depth_solve rules)
   190   by (depth_solve rules)
   191 
   191 
   192 schematic_lemma (in CC) "A:* P:A->* a:A |- ?p : (Pi a:A. P^a)->P^a"
   192 schematic_lemma (in CC) "A:* P:A\<rightarrow>* a:A \<turnstile> ?p : (\<Pi> a:A. P^a)\<rightarrow>P^a"
   193   apply (strip_asms rules)
   193   apply (strip_asms rules)
   194   apply (rule lam_ss)
   194   apply (rule lam_ss)
   195     apply (depth_solve1 rules)
   195     apply (depth_solve1 rules)
   196    prefer 2
   196    prefer 2
   197    apply (depth_solve1 rules)
   197    apply (depth_solve1 rules)
   199   done
   199   done
   200 
   200 
   201 
   201 
   202 subsection {* Some random examples *}
   202 subsection {* Some random examples *}
   203 
   203 
   204 schematic_lemma (in LP2) "A:* c:A f:A->A |-
   204 schematic_lemma (in LP2) "A:* c:A f:A\<rightarrow>A \<turnstile>
   205     Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
   205     \<Lambda> a:A. \<Pi> P:A\<rightarrow>*.P^c \<rightarrow> (\<Pi> x:A. P^x\<rightarrow>P^(f^x)) \<rightarrow> P^a : ?T"
   206   by (depth_solve rules)
   206   by (depth_solve rules)
   207 
   207 
   208 schematic_lemma (in CC) "Lam A:*.Lam c:A. Lam f:A->A.
   208 schematic_lemma (in CC) "\<Lambda> A:*.\<Lambda> c:A. \<Lambda> f:A\<rightarrow>A.
   209     Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
   209     \<Lambda> a:A. \<Pi> P:A\<rightarrow>*.P^c \<rightarrow> (\<Pi> x:A. P^x\<rightarrow>P^(f^x)) \<rightarrow> P^a : ?T"
   210   by (depth_solve rules)
   210   by (depth_solve rules)
   211 
   211 
   212 schematic_lemma (in LP2)
   212 schematic_lemma (in LP2)
   213   "A:* a:A b:A |- ?p: (Pi P:A->*.P^a->P^b) -> (Pi P:A->*.P^b->P^a)"
   213   "A:* a:A b:A \<turnstile> ?p: (\<Pi> P:A\<rightarrow>*.P^a\<rightarrow>P^b) \<rightarrow> (\<Pi> P:A\<rightarrow>*.P^b\<rightarrow>P^a)"
   214   -- {* Symmetry of Leibnitz equality *}
   214   -- {* Symmetry of Leibnitz equality *}
   215   apply (strip_asms rules)
   215   apply (strip_asms rules)
   216   apply (rule lam_ss)
   216   apply (rule lam_ss)
   217     apply (depth_solve1 rules)
   217     apply (depth_solve1 rules)
   218    prefer 2
   218    prefer 2
   219    apply (depth_solve1 rules)
   219    apply (depth_solve1 rules)
   220   apply (erule_tac a = "Lam x:A. Pi Q:A->*.Q^x->Q^a" in pi_elim)
   220   apply (erule_tac a = "\<Lambda> x:A. \<Pi> Q:A\<rightarrow>*.Q^x\<rightarrow>Q^a" in pi_elim)
   221    apply (depth_solve1 rules)
   221    apply (depth_solve1 rules)
   222   apply (unfold beta)
   222   apply (unfold beta)
   223   apply (erule imp_elim)
   223   apply (erule imp_elim)
   224    apply (rule lam_bs)
   224    apply (rule lam_bs)
   225      apply (depth_solve1 rules)
   225      apply (depth_solve1 rules)